added AA_Map; tuned titles
authornipkow
Mon, 11 Jan 2016 20:51:13 +0100
changeset 62130 90a3016a6c12
parent 62129 72d19e588e97
child 62132 09c2a77f91d3
added AA_Map; tuned titles
src/HOL/Data_Structures/AA_Map.thy
src/HOL/Data_Structures/AA_Set.thy
src/HOL/Data_Structures/Brother12_Map.thy
src/HOL/Data_Structures/Brother12_Set.thy
src/HOL/Data_Structures/Tree234_Map.thy
src/HOL/Data_Structures/Tree234_Set.thy
src/HOL/Data_Structures/Tree23_Map.thy
src/HOL/Data_Structures/Tree23_Set.thy
src/HOL/ROOT
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Data_Structures/AA_Map.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -0,0 +1,54 @@
+(* Author: Tobias Nipkow *)
+
+section "AA Implementation of Maps"
+
+theory AA_Map
+imports
+  AA_Set
+  Lookup2
+begin
+
+fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
+"update x y Leaf = Node 1 Leaf (x,y) Leaf" |
+"update x y (Node lv t1 (a,b) t2) =
+  (case cmp x a of
+     LT \<Rightarrow> split (skew (Node lv (update x y t1) (a,b) t2)) |
+     GT \<Rightarrow> split (skew (Node lv t1 (a,b) (update x y t2))) |
+     EQ \<Rightarrow> Node lv t1 (x,y) t2)"
+
+fun delete :: "'a::cmp \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
+"delete _ Leaf = Leaf" |
+"delete x (Node lv l (a,b) r) =
+  (case cmp x a of
+     LT \<Rightarrow> adjust (Node lv (delete x l) (a,b) r) |
+     GT \<Rightarrow> adjust (Node lv l (a,b) (delete x r)) |
+     EQ \<Rightarrow> (if l = Leaf then r
+            else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))"
+
+
+subsection {* Functional Correctness Proofs *}
+
+theorem inorder_update:
+  "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
+by (induct t) (auto simp: upd_list_simps inorder_split inorder_skew)
+
+
+theorem inorder_delete:
+  "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
+by(induction t)
+  (auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits)
+
+interpretation Map_by_Ordered
+where empty = Leaf and lookup = lookup and update = update and delete = delete
+and inorder = inorder and inv = "\<lambda>_. True"
+proof (standard, goal_cases)
+  case 1 show ?case by simp
+next
+  case 2 thus ?case by(simp add: lookup_map_of)
+next
+  case 3 thus ?case by(simp add: inorder_update)
+next
+  case 4 thus ?case by(simp add: inorder_delete)
+qed auto
+
+end
--- a/src/HOL/Data_Structures/AA_Set.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/AA_Set.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,9 +1,10 @@
 (*
 Author: Tobias Nipkow
-Invariants are under development
+
+Added trivial cases to function `adjust' to obviate invariants.
 *)
 
-section \<open>An AA Tree Implementation of Sets\<close>
+section \<open>AA Tree Implementation of Sets\<close>
 
 theory AA_Set
 imports
@@ -16,13 +17,13 @@
 fun lvl :: "'a aa_tree \<Rightarrow> nat" where
 "lvl Leaf = 0" |
 "lvl (Node lv _ _ _) = lv"
-
+(*
 fun invar :: "'a aa_tree \<Rightarrow> bool" where
 "invar Leaf = True" |
 "invar (Node h l a r) =
  (invar l \<and> invar r \<and>
   h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node h lr b rr \<and> h = lvl rr + 1)))"
-
+*)
 fun skew :: "'a aa_tree \<Rightarrow> 'a aa_tree" where
 "skew (Node lva (Node lvb t1 b t2) a t3) =
   (if lva = lvb then Node lva t1 b (Node lva t2 a t3) else Node lva (Node lvb t1 b t2) a t3)" |
@@ -105,8 +106,7 @@
 subsubsection "Proofs for delete"
 
 lemma del_maxD:
-  "\<lbrakk> del_max t = (t',x); t \<noteq> Leaf; sorted(inorder t) \<rbrakk> \<Longrightarrow>
-   inorder t' @ [x] = inorder t"
+  "\<lbrakk> del_max t = (t',x); t \<noteq> Leaf \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
 by(induction t arbitrary: t' rule: del_max.induct)
   (auto simp: sorted_lems split: prod.splits)
 
--- a/src/HOL/Data_Structures/Brother12_Map.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Brother12_Map.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 1-2 Brother Tree Implementation of Maps\<close>
+section \<open>1-2 Brother Tree Implementation of Maps\<close>
 
 theory Brother12_Map
 imports
--- a/src/HOL/Data_Structures/Brother12_Set.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Brother12_Set.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 1-2 Brother Tree Implementation of Sets\<close>
+section \<open>1-2 Brother Tree Implementation of Sets\<close>
 
 theory Brother12_Set
 imports
--- a/src/HOL/Data_Structures/Tree234_Map.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Tree234_Map.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 2-3-4 Tree Implementation of Maps\<close>
+section \<open>2-3-4 Tree Implementation of Maps\<close>
 
 theory Tree234_Map
 imports
--- a/src/HOL/Data_Structures/Tree234_Set.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Tree234_Set.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 2-3-4 Tree Implementation of Sets\<close>
+section \<open>2-3-4 Tree Implementation of Sets\<close>
 
 theory Tree234_Set
 imports
--- a/src/HOL/Data_Structures/Tree23_Map.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Tree23_Map.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 2-3 Tree Implementation of Maps\<close>
+section \<open>2-3 Tree Implementation of Maps\<close>
 
 theory Tree23_Map
 imports
--- a/src/HOL/Data_Structures/Tree23_Set.thy	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/Data_Structures/Tree23_Set.thy	Mon Jan 11 20:51:13 2016 +0100
@@ -1,6 +1,6 @@
 (* Author: Tobias Nipkow *)
 
-section \<open>A 2-3 Tree Implementation of Sets\<close>
+section \<open>2-3 Tree Implementation of Sets\<close>
 
 theory Tree23_Set
 imports
--- a/src/HOL/ROOT	Mon Jan 11 18:27:27 2016 +0100
+++ b/src/HOL/ROOT	Mon Jan 11 20:51:13 2016 +0100
@@ -180,7 +180,7 @@
     Tree23_Map
     Tree234_Map
     Brother12_Map
-    AA_Set
+    AA_Map
     Splay_Map
   document_files "root.tex" "root.bib"