del_max -> split_max
authornipkow
Mon, 23 Apr 2018 08:09:50 +0200
changeset 68023 75130777ece4
parent 68022 c8a506be83bd
child 68024 b5e29bf0aeab
del_max -> split_max
src/HOL/Data_Structures/AA_Map.thy
src/HOL/Data_Structures/AA_Set.thy
src/HOL/Data_Structures/AVL_Map.thy
src/HOL/Data_Structures/AVL_Set.thy
--- a/src/HOL/Data_Structures/AA_Map.thy	Sun Apr 22 21:05:14 2018 +0100
+++ b/src/HOL/Data_Structures/AA_Map.thy	Mon Apr 23 08:09:50 2018 +0200
@@ -23,7 +23,7 @@
      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)))"
+            else let (l',ab') = split_max l in adjust (Node lv l' ab' r)))"
 
 
 subsection "Invariance"
@@ -187,7 +187,7 @@
         by(auto simp: post_del_def invar.simps(2))
     next
       assume "l \<noteq> Leaf" thus ?thesis using equal Node.prems
-        by simp (metis inv_l post_del_adjL post_del_max pre_adj_if_postL)
+        by simp (metis inv_l post_del_adjL post_split_max pre_adj_if_postL)
     qed
   qed
 qed (simp add: post_del_def)
@@ -204,7 +204,7 @@
   inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
-              post_del_max post_delete del_maxD split: prod.splits)
+              post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
--- a/src/HOL/Data_Structures/AA_Set.thy	Sun Apr 22 21:05:14 2018 +0100
+++ b/src/HOL/Data_Structures/AA_Set.thy	Mon Apr 23 08:09:50 2018 +0200
@@ -72,14 +72,14 @@
 text\<open>In the paper, the last case of @{const adjust} is expressed with the help of an
 incorrect auxiliary function \texttt{nlvl}.
 
-Function @{text del_max} below is called \texttt{dellrg} in the paper.
+Function @{text split_max} below is called \texttt{dellrg} in the paper.
 The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest
 element but recurses on the left instead of the right subtree; the invariant
 is not restored.\<close>
 
-fun del_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
-"del_max (Node lv l a Leaf) = (l,a)" |
-"del_max (Node lv l a r) = (let (r',b) = del_max r in (adjust(Node lv l a r'), b))"
+fun split_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where
+"split_max (Node lv l a Leaf) = (l,a)" |
+"split_max (Node lv l a r) = (let (r',b) = split_max r in (adjust(Node lv l a r'), b))"
 
 fun delete :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where
 "delete _ Leaf = Leaf" |
@@ -88,7 +88,7 @@
      LT \<Rightarrow> adjust (Node lv (delete x l) a r) |
      GT \<Rightarrow> adjust (Node lv l a (delete x r)) |
      EQ \<Rightarrow> (if l = Leaf then r
-            else let (l',b) = del_max l in adjust (Node lv l' b r)))"
+            else let (l',b) = split_max l in adjust (Node lv l' b r)))"
 
 fun pre_adjust where
 "pre_adjust (Node lv l a r) = (invar l \<and> invar r \<and>
@@ -397,13 +397,13 @@
 
 declare prod.splits[split]
 
-theorem post_del_max:
- "\<lbrakk> invar t; (t', x) = del_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
-proof (induction t arbitrary: t' rule: del_max.induct)
+theorem post_split_max:
+ "\<lbrakk> invar t; (t', x) = split_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'"
+proof (induction t arbitrary: t' rule: split_max.induct)
   case (2 lv l a lvr rl ra rr)
   let ?r =  "\<langle>lvr, rl, ra, rr\<rangle>"
   let ?t = "\<langle>lv, l, a, ?r\<rangle>"
-  from "2.prems"(2) obtain r' where r': "(r', x) = del_max ?r"
+  from "2.prems"(2) obtain r' where r': "(r', x) = split_max ?r"
     and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto
   from  "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp
   note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> post]
@@ -440,7 +440,7 @@
         by(auto simp: post_del_def invar.simps(2))
     next
       assume "l \<noteq> Leaf" thus ?thesis using equal
-        by simp (metis Node.prems inv_l post_del_adjL post_del_max pre_adj_if_postL)
+        by simp (metis Node.prems inv_l post_del_adjL post_split_max pre_adj_if_postL)
     qed
   qed
 qed (simp add: post_del_def)
@@ -471,16 +471,16 @@
   (auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps
      split: tree.splits)
 
-lemma del_maxD:
-  "\<lbrakk> del_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
-by(induction t arbitrary: t' rule: del_max.induct)
-  (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_del_max split: prod.splits)
+lemma split_maxD:
+  "\<lbrakk> split_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t"
+by(induction t arbitrary: t' rule: split_max.induct)
+  (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_split_max split: prod.splits)
 
 lemma inorder_delete:
   "invar t \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR 
-              post_del_max post_delete del_maxD split: prod.splits)
+              post_split_max post_delete split_maxD split: prod.splits)
 
 interpretation I: Set_by_Ordered
 where empty = Leaf and isin = isin and insert = insert and delete = delete
--- a/src/HOL/Data_Structures/AVL_Map.thy	Sun Apr 22 21:05:14 2018 +0100
+++ b/src/HOL/Data_Structures/AVL_Map.thy	Mon Apr 23 08:09:50 2018 +0200
@@ -34,7 +34,7 @@
   "sorted1(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_balL inorder_balR
-     inorder_del_root inorder_del_maxD split: prod.splits)
+     inorder_del_root inorder_split_maxD split: prod.splits)
 
 interpretation Map_by_Ordered
 where empty = Leaf and lookup = lookup and update = update and delete = delete
--- a/src/HOL/Data_Structures/AVL_Set.thy	Sun Apr 22 21:05:14 2018 +0100
+++ b/src/HOL/Data_Structures/AVL_Set.thy	Mon Apr 23 08:09:50 2018 +0200
@@ -58,16 +58,16 @@
    LT \<Rightarrow> balL (insert x l) a r |
    GT \<Rightarrow> balR l a (insert x r))"
 
-fun del_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
-"del_max (Node _ l a r) =
-  (if r = Leaf then (l,a) else let (r',a') = del_max r in (balL l a r', a'))"
+fun split_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where
+"split_max (Node _ l a r) =
+  (if r = Leaf then (l,a) else let (r',a') = split_max r in (balL l a r', a'))"
 
-lemmas del_max_induct = del_max.induct[case_names Node Leaf]
+lemmas split_max_induct = split_max.induct[case_names Node Leaf]
 
 fun del_root :: "'a avl_tree \<Rightarrow> 'a avl_tree" where
 "del_root (Node h Leaf a r) = r" |
 "del_root (Node h l a Leaf) = l" |
-"del_root (Node h l a r) = (let (l', a') = del_max l in balR l' a' r)"
+"del_root (Node h l a r) = (let (l', a') = split_max l in balR l' a' r)"
 
 lemmas del_root_cases = del_root.cases[case_names Leaf_t Node_Leaf Node_Node]
 
@@ -103,22 +103,22 @@
 
 subsubsection "Proofs for delete"
 
-lemma inorder_del_maxD:
-  "\<lbrakk> del_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
+lemma inorder_split_maxD:
+  "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
    inorder t' @ [a] = inorder t"
-by(induction t arbitrary: t' rule: del_max.induct)
+by(induction t arbitrary: t' rule: split_max.induct)
   (auto simp: inorder_balL split: if_splits prod.splits tree.split)
 
 lemma inorder_del_root:
   "inorder (del_root (Node h l a r)) = inorder l @ inorder r"
 by(cases "Node h l a r" rule: del_root.cases)
-  (auto simp: inorder_balL inorder_balR inorder_del_maxD split: if_splits prod.splits)
+  (auto simp: inorder_balL inorder_balR inorder_split_maxD split: if_splits prod.splits)
 
 theorem inorder_delete:
   "sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
 by(induction t)
   (auto simp: del_list_simps inorder_balL inorder_balR
-    inorder_del_root inorder_del_maxD split: prod.splits)
+    inorder_del_root inorder_split_maxD split: prod.splits)
 
 
 subsubsection "Overall functional correctness"
@@ -301,12 +301,12 @@
 
 subsubsection \<open>Deletion maintains AVL balance\<close>
 
-lemma avl_del_max:
+lemma avl_split_max:
   assumes "avl x" and "x \<noteq> Leaf"
-  shows "avl (fst (del_max x))" "height x = height(fst (del_max x)) \<or>
-         height x = height(fst (del_max x)) + 1"
+  shows "avl (fst (split_max x))" "height x = height(fst (split_max x)) \<or>
+         height x = height(fst (split_max x)) + 1"
 using assms
-proof (induct x rule: del_max_induct)
+proof (induct x rule: split_max_induct)
   case (Node h l a r)
   case 1
   thus ?case using Node
@@ -316,7 +316,7 @@
 next
   case (Node h l a r)
   case 2
-  let ?r' = "fst (del_max r)"
+  let ?r' = "fst (split_max r)"
   from \<open>avl x\<close> Node 2 have "avl l" and "avl r" by simp_all
   thus ?case using Node 2 height_balL[of l ?r' a] height_balL2[of l ?r' a]
     apply (auto split:prod.splits simp del:avl.simps) by arith+
@@ -330,14 +330,14 @@
   case (Node_Node h lh ll ln lr n rh rl rn rr)
   let ?l = "Node lh ll ln lr"
   let ?r = "Node rh rl rn rr"
-  let ?l' = "fst (del_max ?l)"
+  let ?l' = "fst (split_max ?l)"
   from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
   from \<open>avl t\<close> and Node_Node have "avl ?l" by simp
   hence "avl(?l')" "height ?l = height(?l') \<or>
-         height ?l = height(?l') + 1" by (rule avl_del_max,simp)+
+         height ?l = height(?l') + 1" by (rule avl_split_max,simp)+
   with \<open>avl t\<close> Node_Node have "height ?l' = height ?r \<or> height ?l' = height ?r + 1
             \<or> height ?r = height ?l' + 1 \<or> height ?r = height ?l' + 2" by fastforce
-  with \<open>avl ?l'\<close> \<open>avl ?r\<close> have "avl(balR ?l' (snd(del_max ?l)) ?r)"
+  with \<open>avl ?l'\<close> \<open>avl ?r\<close> have "avl(balR ?l' (snd(split_max ?l)) ?r)"
     by (rule avl_balR)
   with Node_Node show ?thesis by (auto split:prod.splits)
 qed simp_all
@@ -350,12 +350,12 @@
   case (Node_Node h lh ll ln lr n rh rl rn rr)
   let ?l = "Node lh ll ln lr"
   let ?r = "Node rh rl rn rr"
-  let ?l' = "fst (del_max ?l)"
-  let ?t' = "balR ?l' (snd(del_max ?l)) ?r"
+  let ?l' = "fst (split_max ?l)"
+  let ?t' = "balR ?l' (snd(split_max ?l)) ?r"
   from \<open>avl t\<close> and Node_Node have "avl ?r" by simp
   from \<open>avl t\<close> and Node_Node have "avl ?l" by simp
-  hence "avl(?l')"  by (rule avl_del_max,simp)
-  have l'_height: "height ?l = height ?l' \<or> height ?l = height ?l' + 1" using \<open>avl ?l\<close> by (intro avl_del_max) auto
+  hence "avl(?l')"  by (rule avl_split_max,simp)
+  have l'_height: "height ?l = height ?l' \<or> height ?l = height ?l' + 1" using \<open>avl ?l\<close> by (intro avl_split_max) auto
   have t_height: "height t = 1 + max (height ?l) (height ?r)" using \<open>avl t\<close> Node_Node by simp
   have "height t = height ?t' \<or> height t = height ?t' + 1" using  \<open>avl t\<close> Node_Node
   proof(cases "height ?r = height ?l' + 2")
@@ -364,7 +364,7 @@
   next
     case True
     show ?thesis
-    proof(cases rule: disjE[OF height_balR[OF True \<open>avl ?l'\<close> \<open>avl ?r\<close>, of "snd (del_max ?l)"]])
+    proof(cases rule: disjE[OF height_balR[OF True \<open>avl ?l'\<close> \<open>avl ?r\<close>, of "snd (split_max ?l)"]])
       case 1
       thus ?thesis using l'_height t_height True by arith
     next