--- a/src/HOL/Data_Structures/RBT.thy Tue May 12 10:24:53 2020 +0200
+++ b/src/HOL/Data_Structures/RBT.thy Tue May 12 10:59:59 2020 +0200
@@ -39,18 +39,18 @@
"baldR (R t1 a (B t2 b t3)) c t4 = R (baliL (paint Red t1) a t2) b (B t3 c t4)" |
"baldR t1 a t2 = R t1 a t2"
-fun app :: "'a rbt \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
-"app Leaf t = t" |
-"app t Leaf = t" |
-"app (R t1 a t2) (R t3 c t4) =
- (case app t2 t3 of
+fun join :: "'a rbt \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
+"join Leaf t = t" |
+"join t Leaf = t" |
+"join (R t1 a t2) (R t3 c t4) =
+ (case join t2 t3 of
R u2 b u3 \<Rightarrow> (R (R t1 a u2) b (R u3 c t4)) |
t23 \<Rightarrow> R t1 a (R t23 c t4))" |
-"app (B t1 a t2) (B t3 c t4) =
- (case app t2 t3 of
+"join (B t1 a t2) (B t3 c t4) =
+ (case join t2 t3 of
R u2 b u3 \<Rightarrow> R (B t1 a u2) b (B u3 c t4) |
t23 \<Rightarrow> baldL t1 a (B t23 c t4))" |
-"app t1 (R t2 a t3) = R (app t1 t2) a t3" |
-"app (R t1 a t2) t3 = R t1 a (app t2 t3)"
+"join t1 (R t2 a t3) = R (join t1 t2) a t3" |
+"join (R t1 a t2) t3 = R t1 a (join t2 t3)"
end
--- a/src/HOL/Data_Structures/RBT_Map.thy Tue May 12 10:24:53 2020 +0200
+++ b/src/HOL/Data_Structures/RBT_Map.thy Tue May 12 10:59:59 2020 +0200
@@ -29,7 +29,7 @@
then baldL (del x l) (a,b) r else R (del x l) (a,b) r |
GT \<Rightarrow> if r \<noteq> Leaf\<and> color r = Black
then baldR l (a,b) (del x r) else R l (a,b) (del x r) |
- EQ \<Rightarrow> app l r)"
+ EQ \<Rightarrow> join l r)"
definition delete :: "'a::linorder \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
"delete x t = paint Black (del x t)"
@@ -49,7 +49,7 @@
lemma inorder_del:
"sorted1(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
by(induction x t rule: del.induct)
- (auto simp: del_list_simps inorder_app inorder_baldL inorder_baldR)
+ (auto simp: del_list_simps inorder_join inorder_baldL inorder_baldR)
lemma inorder_delete:
"sorted1(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
@@ -86,7 +86,7 @@
thus ?case proof (elim disjE)
assume "x = y"
with 2 show ?thesis
- by (cases c) (simp_all add: invh_app invc_app)
+ by (cases c) (simp_all add: invh_join invc_join)
next
assume "x < y"
with 2 show ?thesis
--- a/src/HOL/Data_Structures/RBT_Set.thy Tue May 12 10:24:53 2020 +0200
+++ b/src/HOL/Data_Structures/RBT_Set.thy Tue May 12 10:59:59 2020 +0200
@@ -41,7 +41,7 @@
then baldL (del x l) a r else R (del x l) a r |
GT \<Rightarrow> if r \<noteq> Leaf\<and> color r = Black
then baldR l a (del x r) else R l a (del x r) |
- EQ \<Rightarrow> app l r)"
+ EQ \<Rightarrow> join l r)"
definition delete :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
"delete x t = paint Black (del x t)"
@@ -79,15 +79,15 @@
by(cases "(l,a,r)" rule: baldR.cases)
(auto simp: inorder_baliL inorder_baliR inorder_paint)
-lemma inorder_app:
- "inorder(app l r) = inorder l @ inorder r"
-by(induction l r rule: app.induct)
+lemma inorder_join:
+ "inorder(join l r) = inorder l @ inorder r"
+by(induction l r rule: join.induct)
(auto simp: inorder_baldL inorder_baldR split: tree.split color.split)
lemma inorder_del:
"sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
by(induction x t rule: del.induct)
- (auto simp: del_list_simps inorder_app inorder_baldL inorder_baldR)
+ (auto simp: del_list_simps inorder_join inorder_baldL inorder_baldR)
lemma inorder_delete:
"sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
@@ -225,16 +225,16 @@
lemma invc2_baldR: "\<lbrakk> invc l; invc2 r \<rbrakk> \<Longrightarrow>invc2 (baldR l a r)"
by (induct l a r rule: baldR.induct) (auto simp: invc_baliL paint2 invc2I)
-lemma invh_app:
+lemma invh_join:
"\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk>
- \<Longrightarrow> invh (app l r) \<and> bheight (app l r) = bheight l"
-by (induct l r rule: app.induct)
+ \<Longrightarrow> invh (join l r) \<and> bheight (join l r) = bheight l"
+by (induct l r rule: join.induct)
(auto simp: invh_baldL_Black split: tree.splits color.splits)
-lemma invc_app:
+lemma invc_join:
"\<lbrakk> invc l; invc r \<rbrakk> \<Longrightarrow>
- (color l = Black \<and> color r = Black \<longrightarrow> invc (app l r)) \<and> invc2 (app l r)"
-by (induct l r rule: app.induct)
+ (color l = Black \<and> color r = Black \<longrightarrow> invc (join l r)) \<and> invc2 (join l r)"
+by (induct l r rule: join.induct)
(auto simp: invc_baldL invc2I split: tree.splits color.splits)
text \<open>All in one:\<close>
@@ -253,11 +253,11 @@
by (induct l a r rule: baldR.induct)
(auto simp: inv_baliL invh_paint bheight_baliL bheight_paint_Red paint2 invc2I)
-lemma inv_app:
+lemma inv_join:
"\<lbrakk> invh l; invh r; bheight l = bheight r; invc l; invc r \<rbrakk>
- \<Longrightarrow> invh (app l r) \<and> bheight (app l r) = bheight l
- \<and> invc2 (app l r) \<and> (color l = Black \<and> color r = Black \<longrightarrow> invc (app l r))"
-by (induct l r rule: app.induct)
+ \<Longrightarrow> invh (join l r) \<and> bheight (join l r) = bheight l
+ \<and> invc2 (join l r) \<and> (color l = Black \<and> color r = Black \<longrightarrow> invc (join l r))"
+by (induct l r rule: join.induct)
(auto simp: invh_baldL_Black inv_baldL invc2I split: tree.splits color.splits)
lemma neq_LeafD: "t \<noteq> Leaf \<Longrightarrow> \<exists>l x c r. t = Node l (x,c) r"
@@ -268,7 +268,7 @@
(color t = Red \<and> bheight (del x t) = bheight t \<and> invc (del x t) \<or>
color t = Black \<and> bheight (del x t) = bheight t - 1 \<and> invc2 (del x t))"
by(induct x t rule: del.induct)
- (auto simp: inv_baldL inv_baldR inv_app dest!: neq_LeafD)
+ (auto simp: inv_baldL inv_baldR inv_join dest!: neq_LeafD)
theorem rbt_delete: "rbt t \<Longrightarrow> rbt (delete x t)"
by (metis delete_def rbt_def color_paint_Black inv_del invc2I invh_paint)
--- a/src/HOL/Data_Structures/RBT_Set2.thy Tue May 12 10:24:53 2020 +0200
+++ b/src/HOL/Data_Structures/RBT_Set2.thy Tue May 12 10:59:59 2020 +0200
@@ -6,7 +6,7 @@
imports RBT_Set
begin
-text \<open>This is a conceptually simpler version of deletion. Instead of the tricky \<open>combine\<close>
+text \<open>This is a conceptually simpler version of deletion. Instead of the tricky \<open>join\<close>
function this version follows the standard approach of replacing the deleted element
(in function \<open>del\<close>) by the minimal element in its right subtree.\<close>