src/HOL/Data_Structures/AA_Map.thy
changeset 68413 b56ed5010e69
parent 68023 75130777ece4
child 68431 b294e095f64c
--- a/src/HOL/Data_Structures/AA_Map.thy	Mon Jun 11 08:15:43 2018 +0200
+++ b/src/HOL/Data_Structures/AA_Map.thy	Mon Jun 11 16:29:27 2018 +0200
@@ -9,21 +9,21 @@
 begin
 
 fun update :: "'a::linorder \<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) =
+"update x y Leaf = Node Leaf (x,y) 1 Leaf" |
+"update x y (Node t1 (a,b) lv 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)"
+     LT \<Rightarrow> split (skew (Node (update x y t1) (a,b) lv t2)) |
+     GT \<Rightarrow> split (skew (Node t1 (a,b) lv (update x y t2))) |
+     EQ \<Rightarrow> Node t1 (x,y) lv t2)"
 
 fun delete :: "'a::linorder \<Rightarrow> ('a*'b) aa_tree \<Rightarrow> ('a*'b) aa_tree" where
 "delete _ Leaf = Leaf" |
-"delete x (Node lv l (a,b) r) =
+"delete x (Node l (a,b) lv 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)) |
+     LT \<Rightarrow> adjust (Node (delete x l) (a,b) lv r) |
+     GT \<Rightarrow> adjust (Node l (a,b) lv (delete x r)) |
      EQ \<Rightarrow> (if l = Leaf then r
-            else let (l',ab') = split_max l in adjust (Node lv l' ab' r)))"
+            else let (l',ab') = split_max l in adjust (Node l' ab' lv r)))"
 
 
 subsection "Invariance"
@@ -46,7 +46,7 @@
 
 lemma lvl_update_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(update x y t) = lvl t"
 proof (induction t rule: update.induct)
-  case (2 x y lv t1 a b t2)
+  case (2 x y t1 a b lv t2)
   consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" 
     using less_linear by blast 
   thus ?case proof cases
@@ -64,7 +64,7 @@
 qed simp
 
 lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \<longleftrightarrow>
-  (\<exists>l x r. update a b t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)"
+  (\<exists>l x r. update a b t = Node l x (lvl t + 1) r \<and> lvl l = lvl r)"
 apply(cases t)
 apply(auto simp add: skew_case split_case split: if_splits)
 apply(auto split: tree.splits if_splits)
@@ -72,12 +72,12 @@
 
 lemma invar_update: "invar t \<Longrightarrow> invar(update a b t)"
 proof(induction t)
-  case N: (Node n l xy r)
+  case N: (Node l xy n r)
   hence il: "invar l" and ir: "invar r" by auto
   note iil = N.IH(1)[OF il]
   note iir = N.IH(2)[OF ir]
   obtain x y where [simp]: "xy = (x,y)" by fastforce
-  let ?t = "Node n l xy r"
+  let ?t = "Node l xy n r"
   have "a < x \<or> a = x \<or> x < a" by auto
   moreover
   have ?case if "a < x"
@@ -87,13 +87,13 @@
       by (simp add: skew_invar split_invar del: invar.simps)
   next
     case (Incr)
-    then obtain t1 w t2 where ial[simp]: "update a b l = Node n t1 w t2"
+    then obtain t1 w t2 where ial[simp]: "update a b l = Node t1 w n t2"
       using N.prems by (auto simp: lvl_Suc_iff)
     have l12: "lvl t1 = lvl t2"
       by (metis Incr(1) ial lvl_update_incr_iff tree.inject)
-    have "update a b ?t = split(skew(Node n (update a b l) xy r))"
+    have "update a b ?t = split(skew(Node (update a b l) xy n r))"
       by(simp add: \<open>a<x\<close>)
-    also have "skew(Node n (update a b l) xy r) = Node n t1 w (Node n t2 xy r)"
+    also have "skew(Node (update a b l) xy n r) = Node t1 w n (Node t2 xy n r)"
       by(simp)
     also have "invar(split \<dots>)"
     proof (cases r)
@@ -101,7 +101,7 @@
       hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff)
       thus ?thesis using Leaf ial by simp
     next
-      case [simp]: (Node m t3 y t4)
+      case [simp]: (Node t3 y m t4)
       show ?thesis (*using N(3) iil l12 by(auto)*)
       proof cases
         assume "m = n" thus ?thesis using N(3) iil by(auto)
@@ -118,14 +118,14 @@
     thus ?case
     proof
       assume 0: "n = lvl r"
-      have "update a b ?t = split(skew(Node n l xy (update a b r)))"
+      have "update a b ?t = split(skew(Node l xy n (update a b r)))"
         using \<open>a>x\<close> by(auto)
-      also have "skew(Node n l xy (update a b r)) = Node n l xy (update a b r)"
+      also have "skew(Node l xy n (update a b r)) = Node l xy n (update a b r)"
         using N.prems by(simp add: skew_case split: tree.split)
       also have "invar(split \<dots>)"
       proof -
         from lvl_update_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a b]
-        obtain t1 p t2 where iar: "update a b r = Node n t1 p t2"
+        obtain t1 p t2 where iar: "update a b r = Node t1 p n t2"
           using N.prems 0 by (auto simp: lvl_Suc_iff)
         from N.prems iar 0 iir
         show ?thesis by (auto simp: split_case split: tree.splits)
@@ -157,12 +157,12 @@
 
 theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)"
 proof (induction t)
-  case (Node lv l ab r)
+  case (Node l ab lv r)
 
   obtain a b where [simp]: "ab = (a,b)" by fastforce
 
   let ?l' = "delete x l" and ?r' = "delete x r"
-  let ?t = "Node lv l ab r" let ?t' = "delete x ?t"
+  let ?t = "Node l ab lv r" let ?t' = "delete x ?t"
 
   from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto)