# HG changeset patch # User nipkow # Date 1528742751 -7200 # Node ID d74ba11680d4ebac857d14cd33f5e92283270c7b # Parent b001bef9aa39981e85655abe790af8c53d534295 tuned def. of del and proved preservation of rbt (finally) diff -r b001bef9aa39 -r d74ba11680d4 src/HOL/Data_Structures/RBT_Map.thy --- a/src/HOL/Data_Structures/RBT_Map.thy Mon Jun 11 16:29:38 2018 +0200 +++ b/src/HOL/Data_Structures/RBT_Map.thy Mon Jun 11 20:45:51 2018 +0200 @@ -22,19 +22,14 @@ definition update :: "'a::linorder \ 'b \ ('a*'b) rbt \ ('a*'b) rbt" where "update x y t = paint Black (upd x y t)" -fun del :: "'a::linorder \ ('a*'b)rbt \ ('a*'b)rbt" -and delL :: "'a::linorder \ ('a*'b)rbt \ 'a*'b \ ('a*'b)rbt \ ('a*'b)rbt" -and delR :: "'a::linorder \ ('a*'b)rbt \ 'a*'b \ ('a*'b)rbt \ ('a*'b)rbt" -where +fun del :: "'a::linorder \ ('a*'b)rbt \ ('a*'b)rbt" where "del x Leaf = Leaf" | -"del x (Node t1 (a,b) c t2) = (case cmp x a of - LT \ delL x t1 (a,b) t2 | - GT \ delR x t1 (a,b) t2 | - EQ \ combine t1 t2)" | -"delL x (B t1 a t2) b t3 = baldL (del x (B t1 a t2)) b t3" | -"delL x t1 a t2 = R (del x t1) a t2" | -"delR x t1 a (B t2 b t3) = baldR t1 a (del x (B t2 b t3))" | -"delR x t1 a t2 = R t1 a (del x t2)" +"del x (Node l (a,b) c r) = (case cmp x a of + LT \ if l \ Leaf \ color l = Black + 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)" definition delete :: "'a::linorder \ ('a*'b) rbt \ ('a*'b) rbt" where "delete x t = paint Black (del x t)" @@ -52,21 +47,66 @@ by(simp add: update_def inorder_upd inorder_paint) lemma inorder_del: - "sorted1(inorder t1) \ inorder(del x t1) = del_list x (inorder t1)" and - "sorted1(inorder t1) \ inorder(delL x t1 a t2) = - del_list x (inorder t1) @ a # inorder t2" and - "sorted1(inorder t2) \ inorder(delR x t1 a t2) = - inorder t1 @ a # del_list x (inorder t2)" -by(induction x t1 and x t1 a t2 and x t1 a t2 rule: del_delL_delR.induct) + "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) lemma inorder_delete: "sorted1(inorder t) \ inorder(delete x t) = del_list x (inorder t)" by(simp add: delete_def inorder_del inorder_paint) + +subsection \Structural invariants\ + +subsubsection \Update\ + +lemma invc_upd: assumes "invc t" + shows "color t = Black \ invc (upd x y t)" "invc2 (upd x y t)" +using assms +by (induct x y t rule: upd.induct) (auto simp: invc_baliL invc_baliR invc2I) + +lemma invh_upd: assumes "invh t" + shows "invh (upd x y t)" "bheight (upd x y t) = bheight t" +using assms +by(induct x y t rule: upd.induct) + (auto simp: invh_baliL invh_baliR bheight_baliL bheight_baliR) + +theorem rbt_update: "rbt t \ rbt (update x y t)" +by (simp add: invc_upd(2) invh_upd(1) color_paint_Black invc_paint_Black invh_paint + rbt_def update_def) + + +subsubsection \Deletion\ + +lemma del_invc_invh: "invh t \ invc t \ invh (del x t) \ + (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))" +proof (induct x t rule: del.induct) +case (2 x _ y _ c) + have "x = y \ x < y \ x > y" by auto + thus ?case proof (elim disjE) + assume "x = y" + with 2 show ?thesis + by (cases c) (simp_all add: invh_combine invc_combine) + next + assume "x < y" + with 2 show ?thesis + by(cases c) + (auto simp: invh_baldL_invc invc_baldL invc2_baldL dest: neq_LeafD) + next + assume "y < x" + with 2 show ?thesis + by(cases c) + (auto simp: invh_baldR_invc invc_baldR invc2_baldR dest: neq_LeafD) + qed +qed auto + +theorem rbt_delete: "rbt t \ rbt (delete k t)" +by (metis delete_def rbt_def color_paint_Black del_invc_invh invc_paint_Black invc2I invh_paint) + interpretation Map_by_Ordered where empty = Leaf and lookup = lookup and update = update and delete = delete -and inorder = inorder and inv = "\_. True" +and inorder = inorder and inv = rbt proof (standard, goal_cases) case 1 show ?case by simp next @@ -75,6 +115,12 @@ case 3 thus ?case by(simp add: inorder_update) next case 4 thus ?case by(simp add: inorder_delete) -qed auto +next + case 5 thus ?case by (simp add: rbt_Leaf) +next + case 6 thus ?case by (simp add: rbt_update) +next + case 7 thus ?case by (simp add: rbt_delete) +qed end