# HG changeset patch # User nipkow # Date 1676475892 -3600 # Node ID 40b23105a322b729065987db3c7370733477962e # Parent bc43f86c9598150879a6ff65ee0592d8dc1a3f2f# Parent d1ca1e587a8ed6a3add2cb84f3f0b9400038e690 merged diff -r bc43f86c9598 -r 40b23105a322 src/HOL/Data_Structures/RBT_Map.thy --- a/src/HOL/Data_Structures/RBT_Map.thy Wed Feb 15 10:56:23 2023 +0100 +++ b/src/HOL/Data_Structures/RBT_Map.thy Wed Feb 15 16:44:52 2023 +0100 @@ -24,11 +24,11 @@ fun del :: "'a::linorder \ ('a*'b)rbt \ ('a*'b)rbt" where "del x Leaf = Leaf" | -"del x (Node l ((a,b), c) r) = (case cmp x a of +"del x (Node l (ab, _) r) = (case cmp x (fst ab) of LT \ if l \ Leaf \ color l = Black - then baldL (del x l) (a,b) r else R (del x l) (a,b) r | + then baldL (del x l) ab r else R (del x l) ab r | GT \ if r \ Leaf\ color r = Black - then baldR l (a,b) (del x r) else R l (a,b) (del x r) | + then baldR l ab (del x r) else R l ab (del x r) | EQ \ join l r)" definition delete :: "'a::linorder \ ('a*'b) rbt \ ('a*'b) rbt" where @@ -46,10 +46,14 @@ "sorted1(inorder t) \ inorder(update x y t) = upd_list x y (inorder t)" by(simp add: update_def inorder_upd inorder_paint) +(* This lemma became necessary below when \del\ was converted from pattern-matching to \fst\ *) +lemma del_list_id: "\ab\set ps. y < fst ab \ x \ y \ del_list x ps = ps" +by(rule del_list_idem) auto + 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_join inorder_baldL inorder_baldR) + (auto simp: del_list_simps del_list_id inorder_join inorder_baldL inorder_baldR) lemma inorder_delete: "sorted1(inorder t) \ inorder(delete x t) = del_list x (inorder t)" @@ -81,19 +85,19 @@ (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 +case (2 x _ ab c) + have "x = fst ab \ x < fst ab \ x > fst ab" by auto thus ?case proof (elim disjE) - assume "x = y" + assume "x = fst ab" with 2 show ?thesis by (cases c) (simp_all add: invh_join invc_join) next - assume "x < y" + assume "x < fst ab" with 2 show ?thesis by(cases c) (auto simp: invh_baldL_invc invc_baldL invc2_baldL dest: neq_LeafD) next - assume "y < x" + assume "fst ab < x" with 2 show ?thesis by(cases c) (auto simp: invh_baldR_invc invc_baldR invc2_baldR dest: neq_LeafD)