# HG changeset patch # User nipkow # Date 1589273999 -7200 # Node ID 7a997ead54b000a07d6b71012fe19cc7b187ae25 # Parent 6f2663df83747eb67886ca5cd9043b41a90d9da9 "app" -> "join" for RBTs diff -r 6f2663df8374 -r 7a997ead54b0 src/HOL/Data_Structures/RBT.thy --- 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 \ 'a rbt \ '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 \ 'a rbt \ '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 \ (R (R t1 a u2) b (R u3 c t4)) | t23 \ 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 \ R (B t1 a u2) b (B u3 c t4) | t23 \ 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 diff -r 6f2663df8374 -r 7a997ead54b0 src/HOL/Data_Structures/RBT_Map.thy --- 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 \ if r \ Leaf\ color r = Black then baldR l (a,b) (del x r) else R l (a,b) (del x r) | - EQ \ app l r)" + EQ \ join l r)" definition delete :: "'a::linorder \ ('a*'b) rbt \ ('a*'b) rbt" where "delete x t = paint Black (del x t)" @@ -49,7 +49,7 @@ lemma inorder_del: "sorted1(inorder t) \ 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) \ 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 diff -r 6f2663df8374 -r 7a997ead54b0 src/HOL/Data_Structures/RBT_Set.thy --- 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 \ if r \ Leaf\ color r = Black then baldR l a (del x r) else R l a (del x r) | - EQ \ app l r)" + EQ \ join l r)" definition delete :: "'a::linorder \ 'a rbt \ '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) \ 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) \ inorder(delete x t) = del_list x (inorder t)" @@ -225,16 +225,16 @@ lemma invc2_baldR: "\ invc l; invc2 r \ \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: "\ invh l; invh r; bheight l = bheight r \ - \ invh (app l r) \ bheight (app l r) = bheight l" -by (induct l r rule: app.induct) + \ invh (join l r) \ 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: "\ invc l; invc r \ \ - (color l = Black \ color r = Black \ invc (app l r)) \ invc2 (app l r)" -by (induct l r rule: app.induct) + (color l = Black \ color r = Black \ invc (join l r)) \ invc2 (join l r)" +by (induct l r rule: join.induct) (auto simp: invc_baldL invc2I split: tree.splits color.splits) text \All in one:\ @@ -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: "\ invh l; invh r; bheight l = bheight r; invc l; invc r \ - \ invh (app l r) \ bheight (app l r) = bheight l - \ invc2 (app l r) \ (color l = Black \ color r = Black \ invc (app l r))" -by (induct l r rule: app.induct) + \ invh (join l r) \ bheight (join l r) = bheight l + \ invc2 (join l r) \ (color l = Black \ color r = Black \ 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 \ Leaf \ \l x c r. t = Node l (x,c) r" @@ -268,7 +268,7 @@ (color t = Red \ bheight (del x t) = bheight t \ invc (del x t) \ color t = Black \ bheight (del x t) = bheight t - 1 \ 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 \ rbt (delete x t)" by (metis delete_def rbt_def color_paint_Black inv_del invc2I invh_paint) diff -r 6f2663df8374 -r 7a997ead54b0 src/HOL/Data_Structures/RBT_Set2.thy --- 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 \This is a conceptually simpler version of deletion. Instead of the tricky \combine\ +text \This is a conceptually simpler version of deletion. Instead of the tricky \join\ function this version follows the standard approach of replacing the deleted element (in function \del\) by the minimal element in its right subtree.\