new theory
authornipkow
Wed, 18 Nov 2020 19:38:19 +0100
changeset 72642 d152890dd17e
parent 72641 4eea17b3ac58
child 72645 f8cc3153ac77
new theory
src/HOL/Data_Structures/Queue_2Lists.thy
src/HOL/Data_Structures/Reverse.thy
--- a/src/HOL/Data_Structures/Queue_2Lists.thy	Wed Nov 18 13:44:34 2020 +0100
+++ b/src/HOL/Data_Structures/Queue_2Lists.thy	Wed Nov 18 19:38:19 2020 +0100
@@ -3,7 +3,9 @@
 section \<open>Queue Implementation via 2 Lists\<close>
 
 theory Queue_2Lists
-imports Queue_Spec
+imports
+  Queue_Spec
+  Reverse
 begin
 
 text \<open>Definitions:\<close>
@@ -11,7 +13,7 @@
 type_synonym 'a queue = "'a list \<times> 'a list"
 
 fun norm :: "'a queue \<Rightarrow> '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 \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
 "enq a (fs,rs) = norm(fs, a # rs)"
@@ -42,7 +44,7 @@
 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
@@ -58,7 +60,7 @@
 text \<open>Running times:\<close>
 
 fun T_norm :: "'a queue \<Rightarrow> 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 \<Rightarrow> 'a queue \<Rightarrow> nat" where
 "T_enq a (fs,rs) = T_norm(fs, a # rs) + 1"
@@ -77,10 +79,10 @@
 fun \<Phi> :: "'a queue \<Rightarrow> nat" where
 "\<Phi>(fs,rs) = length rs"
 
-lemma a_enq: "T_enq a (fs,rs) + \<Phi>(enq a (fs,rs)) - \<Phi>(fs,rs) \<le> 3"
-by(auto)
+lemma a_enq: "T_enq a (fs,rs) + \<Phi>(enq a (fs,rs)) - \<Phi>(fs,rs) \<le> 4"
+by(auto simp: T_itrev)
 
-lemma a_deq: "T_deq (fs,rs) + \<Phi>(deq (fs,rs)) - \<Phi>(fs,rs) \<le> 2"
-by(auto)
+lemma a_deq: "T_deq (fs,rs) + \<Phi>(deq (fs,rs)) - \<Phi>(fs,rs) \<le> 3"
+by(auto simp: T_itrev)
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Data_Structures/Reverse.thy	Wed Nov 18 19:38:19 2020 +0100
@@ -0,0 +1,36 @@
+theory Reverse
+imports Main
+begin
+
+fun T_append :: "'a list \<Rightarrow> 'a list \<Rightarrow> nat" where
+"T_append [] ys = 1" |
+"T_append (x#xs) ys = T_append xs ys + 1"
+
+fun T_rev :: "'a list \<Rightarrow> 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 \<le> (length xs + 1)^2"
+by(induction xs) (auto simp: T_append power2_eq_square)
+
+fun itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> '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 \<Rightarrow> 'a list \<Rightarrow> 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