# HG changeset patch # User nipkow # Date 1566296363 -7200 # Node ID f7c1f585edeb2f00b538d4c08e8eaa4fcaa25c8e # Parent 17909568216ef9460accbb8ff8f22b380a424efe tuned diff -r 17909568216e -r f7c1f585edeb src/HOL/Data_Structures/Set2_Join_RBT.thy --- a/src/HOL/Data_Structures/Set2_Join_RBT.thy Tue Aug 20 09:48:22 2019 +0200 +++ b/src/HOL/Data_Structures/Set2_Join_RBT.thy Tue Aug 20 12:19:23 2019 +0200 @@ -32,7 +32,7 @@ B l' x' r' \ baliR l' x' (joinR r' x r) | R l' x' r' \ R l' x' (joinR r' x r))" -fun join :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where +definition join :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where "join l x r = (if bheight l > bheight r then paint Black (joinR l x r) @@ -102,7 +102,7 @@ (* unused *) lemma rbt_join: "\ invc l; invh l; invc r; invh r \ \ rbt(join l x r)" by(simp add: invc2_joinL invc2_joinR invc_paint_Black invh_joinL invh_joinR invh_paint rbt_def - color_paint_Black) + color_paint_Black join_def) text \To make sure the the black height is not increased unnecessarily:\ @@ -112,7 +112,7 @@ lemma "\ rbt l; rbt r \ \ bheight(join l x r) \ max (bheight l) (bheight r) + 1" using bheight_paint_Black[of "joinL l x r"] bheight_paint_Black[of "joinR l x r"] bheight_joinL[of l r x] bheight_joinR[of l r x] -by(auto simp: max_def rbt_def) +by(auto simp: max_def rbt_def join_def) subsubsection "Inorder properties" @@ -133,7 +133,8 @@ qed lemma "inorder(join l x r) = inorder l @ x # inorder r" -by(auto simp: inorder_joinL inorder_joinR inorder_paint split!: tree.splits color.splits if_splits +by(auto simp: inorder_joinL inorder_joinR inorder_paint join_def + split!: tree.splits color.splits if_splits dest!: arg_cong[where f = inorder]) @@ -165,7 +166,7 @@ by (cases t) auto lemma set_join: "set_tree (join l x r) = set_tree l \ {x} \ set_tree r" -by(simp add: set_joinL set_joinR set_paint) +by(simp add: set_joinL set_joinR set_paint join_def) lemma bst_baliL: "\bst l; bst r; \x\set_tree l. x < a; \x\set_tree r. a < x\ @@ -202,10 +203,10 @@ lemma bst_join: "bst (Node l a n r) \ bst (join l a r)" -by(auto simp: bst_paint bst_joinL bst_joinR) +by(auto simp: bst_paint bst_joinL bst_joinR join_def) lemma inv_join: "\ invc l; invh l; invc r; invh r \ \ invc(join l x r) \ invh(join l x r)" -by (simp add: invc2_joinL invc2_joinR invc_paint_Black invh_joinL invh_joinR invh_paint) +by (simp add: invc2_joinL invc2_joinR invc_paint_Black invh_joinL invh_joinR invh_paint join_def) subsubsection "Interpretation of \<^locale>\Set2_Join\ with Red-Black Tree" @@ -216,11 +217,11 @@ proof (standard, goal_cases) case 1 show ?case by (rule set_join) next - case 2 thus ?case by (simp add: bst_join del: join.simps) + case 2 thus ?case by (simp add: bst_join) next case 3 show ?case by simp next - case 4 thus ?case by (simp add: inv_join del: join.simps) + case 4 thus ?case by (simp add: inv_join) next case 5 thus ?case by simp qed diff -r 17909568216e -r f7c1f585edeb src/HOL/Data_Structures/Set_Specs.thy --- a/src/HOL/Data_Structures/Set_Specs.thy Tue Aug 20 09:48:22 2019 +0200 +++ b/src/HOL/Data_Structures/Set_Specs.thy Tue Aug 20 12:19:23 2019 +0200 @@ -17,7 +17,7 @@ fixes invar :: "'s \ bool" assumes set_empty: "set empty = {}" assumes set_isin: "invar s \ isin s x = (x \ set s)" -assumes set_insert: "invar s \ set(insert x s) = Set.insert x (set s)" +assumes set_insert: "invar s \ set(insert x s) = set s \ {x}" assumes set_delete: "invar s \ set(delete x s) = set s - {x}" assumes invar_empty: "invar empty" assumes invar_insert: "invar s \ invar(insert x s)"