# HG changeset patch # User nipkow # Date 1582303916 -3600 # Node ID a31a9da4369460e87007e79d42819de839e2fdcb # Parent ed8d50969995f4a9c6361f43555a1f881e7222cd tuned deletion diff -r ed8d50969995 -r a31a9da43694 src/HOL/Data_Structures/RBT.thy --- a/src/HOL/Data_Structures/RBT.thy Thu Feb 20 09:05:19 2020 +0100 +++ b/src/HOL/Data_Structures/RBT.thy Fri Feb 21 17:51:56 2020 +0100 @@ -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 combine :: "'a rbt \ 'a rbt \ 'a rbt" where -"combine Leaf t = t" | -"combine t Leaf = t" | -"combine (R t1 a t2) (R t3 c t4) = - (case combine t2 t3 of +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 R u2 b u3 \ (R (R t1 a u2) b (R u3 c t4)) | t23 \ R t1 a (R t23 c t4))" | -"combine (B t1 a t2) (B t3 c t4) = - (case combine t2 t3 of +"app (B t1 a t2) (B t3 c t4) = + (case app 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))" | -"combine t1 (R t2 a t3) = R (combine t1 t2) a t3" | -"combine (R t1 a t2) t3 = R t1 a (combine t2 t3)" +"app t1 (R t2 a t3) = R (app t1 t2) a t3" | +"app (R t1 a t2) t3 = R t1 a (app t2 t3)" end diff -r ed8d50969995 -r a31a9da43694 src/HOL/Data_Structures/RBT_Map.thy --- a/src/HOL/Data_Structures/RBT_Map.thy Thu Feb 20 09:05:19 2020 +0100 +++ b/src/HOL/Data_Structures/RBT_Map.thy Fri Feb 21 17:51:56 2020 +0100 @@ -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 \ combine l r)" + EQ \ app 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_combine inorder_baldL inorder_baldR) + (auto simp: del_list_simps inorder_app 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_combine invc_combine) + by (cases c) (simp_all add: invh_app invc_app) next assume "x < y" with 2 show ?thesis diff -r ed8d50969995 -r a31a9da43694 src/HOL/Data_Structures/RBT_Set.thy --- a/src/HOL/Data_Structures/RBT_Set.thy Thu Feb 20 09:05:19 2020 +0100 +++ b/src/HOL/Data_Structures/RBT_Set.thy Fri Feb 21 17:51:56 2020 +0100 @@ -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 \ combine l r)" + EQ \ app 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_combine: - "inorder(combine l r) = inorder l @ inorder r" -by(induction l r rule: combine.induct) +lemma inorder_app: + "inorder(app l r) = inorder l @ inorder r" +by(induction l r rule: app.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_combine inorder_baldL inorder_baldR) + (auto simp: del_list_simps inorder_app 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_combine: +lemma invh_app: "\ invh l; invh r; bheight l = bheight r \ - \ invh (combine l r) \ bheight (combine l r) = bheight l" -by (induct l r rule: combine.induct) + \ invh (app l r) \ bheight (app l r) = bheight l" +by (induct l r rule: app.induct) (auto simp: invh_baldL_Black split: tree.splits color.splits) -lemma invc_combine: +lemma invc_app: "\ invc l; invc r \ \ - (color l = Black \ color r = Black \ invc (combine l r)) \ invc2 (combine l r)" -by (induct l r rule: combine.induct) + (color l = Black \ color r = Black \ invc (app l r)) \ invc2 (app l r)" +by (induct l r rule: app.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_combine: +lemma inv_app: "\ invh l; invh r; bheight l = bheight r; invc l; invc r \ - \ invh (combine l r) \ bheight (combine l r) = bheight l - \ invc2 (combine l r) \ (color l = Black \ color r = Black \ invc (combine l r))" -by (induct l r rule: combine.induct) + \ 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) (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_combine dest!: neq_LeafD) + (auto simp: inv_baldL inv_baldR inv_app 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 ed8d50969995 -r a31a9da43694 src/HOL/Data_Structures/Tree_Set.thy --- a/src/HOL/Data_Structures/Tree_Set.thy Thu Feb 20 09:05:19 2020 +0100 +++ b/src/HOL/Data_Structures/Tree_Set.thy Fri Feb 21 17:51:56 2020 +0100 @@ -30,6 +30,8 @@ EQ \ Node l a r | GT \ Node l a (insert x r))" +text \Deletion by replacing:\ + fun split_min :: "'a tree \ 'a * 'a tree" where "split_min (Node l a r) = (if l = Leaf then (a,r) else let (x,l') = split_min l in (x, Node l' a r))" @@ -42,6 +44,24 @@ GT \ Node l a (delete x r) | EQ \ if r = Leaf then l else let (a',r') = split_min r in Node l a' r')" +text \Deletion by appending:\ + +fun app :: "('a::linorder)tree \ 'a tree \ 'a tree" where +"app t Leaf = t" | +"app Leaf t = t" | +"app (Node t1 a t2) (Node t3 b t4) = + (case app t2 t3 of + Leaf \ Node t1 a (Node Leaf b t4) | + Node u2 x u3 \ Node (Node t1 a u2) x (Node u3 b t4))" + +fun delete2 :: "'a::linorder \ 'a tree \ 'a tree" where +"delete2 x Leaf = Leaf" | +"delete2 x (Node l a r) = + (case cmp x a of + LT \ Node (delete2 x l) a r | + GT \ Node l a (delete2 x r) | + EQ \ app l r)" + subsection "Functional Correctness Proofs" @@ -75,4 +95,25 @@ case 4 thus ?case by(simp add: inorder_delete) qed (rule TrueI)+ +lemma inorder_app: + "inorder(app l r) = inorder l @ inorder r" +by(induction l r rule: app.induct) (auto split: tree.split) + +lemma inorder_delete2: + "sorted(inorder t) \ inorder(delete2 x t) = del_list x (inorder t)" +by(induction t) (auto simp: inorder_app del_list_simps) + +interpretation S2: Set_by_Ordered +where empty = empty and isin = isin and insert = insert and delete = delete2 +and inorder = inorder and inv = "\_. True" +proof (standard, goal_cases) + case 1 show ?case by (simp add: empty_def) +next + case 2 thus ?case by(simp add: isin_set) +next + case 3 thus ?case by(simp add: inorder_insert) +next + case 4 thus ?case by(simp add: inorder_delete2) +qed (rule TrueI)+ + end