diff --git a/src/HOL/Data_Structures/Queue_2Lists.thy b/src/HOL/Data_Structures/Queue_2Lists.thy --- a/src/HOL/Data_Structures/Queue_2Lists.thy +++ b/src/HOL/Data_Structures/Queue_2Lists.thy @@ -1,86 +1,88 @@ (* Author: Tobias Nipkow *) section \Queue Implementation via 2 Lists\ theory Queue_2Lists -imports Queue_Spec +imports + Queue_Spec + Reverse begin text \Definitions:\ type_synonym 'a queue = "'a list \ 'a list" fun norm :: "'a queue \ 'a queue" where -"norm (fs,rs) = (if fs = [] then (rev rs, []) else (fs,rs))" +"norm (fs,rs) = (if fs = [] then (itrev rs [], []) else (fs,rs))" fun enq :: "'a \ 'a queue \ 'a queue" where "enq a (fs,rs) = norm(fs, a # rs)" fun deq :: "'a queue \ 'a queue" where "deq (fs,rs) = (if fs = [] then (fs,rs) else norm(tl fs,rs))" fun first :: "'a queue \ 'a" where "first (a # fs,rs) = a" fun is_empty :: "'a queue \ bool" where "is_empty (fs,rs) = (fs = [])" fun list :: "'a queue \ 'a list" where "list (fs,rs) = fs @ rev rs" fun invar :: "'a queue \ bool" where "invar (fs,rs) = (fs = [] \ rs = [])" text \Implementation correctness:\ interpretation Queue where empty = "([],[])" and enq = enq and deq = deq and first = first and is_empty = is_empty and list = list and invar = invar proof (standard, goal_cases) case 1 show ?case by (simp) next case (2 q) thus ?case by(cases q) (simp) next - case (3 q) thus ?case by(cases q) (simp) + case (3 q) thus ?case by(cases q) (simp add: itrev_Nil) next case (4 q) thus ?case by(cases q) (auto simp: neq_Nil_conv) next case (5 q) thus ?case by(cases q) (auto) next case 6 show ?case by(simp) next case (7 q) thus ?case by(cases q) (simp) next case (8 q) thus ?case by(cases q) (simp) qed text \Running times:\ fun T_norm :: "'a queue \ nat" where -"T_norm (fs,rs) = (if fs = [] then length rs else 0) + 1" +"T_norm (fs,rs) = (if fs = [] then T_itrev rs [] else 0) + 1" fun T_enq :: "'a \ 'a queue \ nat" where "T_enq a (fs,rs) = T_norm(fs, a # rs) + 1" fun T_deq :: "'a queue \ nat" where "T_deq (fs,rs) = (if fs = [] then 0 else T_norm(tl fs,rs)) + 1" fun T_first :: "'a queue \ nat" where "T_first (a # fs,rs) = 1" fun T_is_empty :: "'a queue \ nat" where "T_is_empty (fs,rs) = 1" text \Amortized running times:\ fun \ :: "'a queue \ nat" where "\(fs,rs) = length rs" -lemma a_enq: "T_enq a (fs,rs) + \(enq a (fs,rs)) - \(fs,rs) \ 3" -by(auto) +lemma a_enq: "T_enq a (fs,rs) + \(enq a (fs,rs)) - \(fs,rs) \ 4" +by(auto simp: T_itrev) -lemma a_deq: "T_deq (fs,rs) + \(deq (fs,rs)) - \(fs,rs) \ 2" -by(auto) +lemma a_deq: "T_deq (fs,rs) + \(deq (fs,rs)) - \(fs,rs) \ 3" +by(auto simp: T_itrev) end diff --git a/src/HOL/Data_Structures/Reverse.thy b/src/HOL/Data_Structures/Reverse.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Data_Structures/Reverse.thy @@ -0,0 +1,36 @@ +theory Reverse +imports Main +begin + +fun T_append :: "'a list \ 'a list \ nat" where +"T_append [] ys = 1" | +"T_append (x#xs) ys = T_append xs ys + 1" + +fun T_rev :: "'a list \ nat" where +"T_rev [] = 1" | +"T_rev (x#xs) = T_rev xs + T_append (rev xs) [x] + 1" + +lemma T_append: "T_append xs ys = length xs + 1" +by(induction xs) auto + +lemma T_rev: "T_rev xs \ (length xs + 1)^2" +by(induction xs) (auto simp: T_append power2_eq_square) + +fun itrev :: "'a list \ 'a list \ 'a list" where +"itrev [] ys = ys" | +"itrev (x#xs) ys = itrev xs (x # ys)" + +lemma itrev: "itrev xs ys = rev xs @ ys" +by(induction xs arbitrary: ys) auto + +lemma itrev_Nil: "itrev xs [] = rev xs" +by(simp add: itrev) + +fun T_itrev :: "'a list \ 'a list \ nat" where +"T_itrev [] ys = 1" | +"T_itrev (x#xs) ys = T_itrev xs (x # ys) + 1" + +lemma T_itrev: "T_itrev xs ys = length xs + 1" +by(induction xs arbitrary: ys) auto + +end \ No newline at end of file