# HG changeset patch # User nipkow # Date 1523182440 -7200 # Node ID f13796496e82379b75be180ffb2d3f3cdc01cae0 # Parent aaa31cd0caef7904aa21b16e625f704d86be0ba2 Added binary set operations with join-based implementation diff -r aaa31cd0caef -r f13796496e82 src/HOL/Data_Structures/Set2_BST2_Join.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Set2_BST2_Join.thy Sun Apr 08 12:14:00 2018 +0200 @@ -0,0 +1,370 @@ +(* Author: Tobias Nipkow *) + +section "Join-Based BST2 Implementation of Sets" + +theory Set2_BST2_Join +imports + Isin2 +begin + +text \Based on theory @{theory Tree2}, otherwise same as theory @{file "Set2_BST_Join.thy"}.\ + +text \This theory implements the set operations \insert\, \delete\, +\union\, \inter\section and \diff\erence. The implementation is based on binary search trees. +All operations are reduced to a single operation \join l x r\ that joins two BSTs \l\ and \r\ +and an element \x\ such that \l < x < r\. + +The theory is based on theory @{theory Tree2} where nodes have an additional field. +This field is ignored here but it means that this theory can be instantiated +with red-black trees (see theory @{file "Set2_BST2_Join_RBT.thy"}) and other balanced trees. +This approach is very concrete and fixes the type of trees. +Alternatively, one could assume some abstract type @{typ 't} of trees with suitable decomposition +and recursion operators on it.\ + +fun set_tree :: "('a,'b) tree \ 'a set" where +"set_tree Leaf = {}" | +"set_tree (Node _ l x r) = Set.insert x (set_tree l \ set_tree r)" + +fun bst :: "('a::linorder,'b) tree \ bool" where +"bst Leaf = True" | +"bst (Node _ l a r) = (bst l \ bst r \ (\x \ set_tree l. x < a) \ (\x \ set_tree r. a < x))" + +lemma isin_iff: "bst t \ isin t x \ x \ set_tree t" +by(induction t) auto + +locale Set2_BST2_Join = +fixes join :: "('a::linorder,'b) tree \ 'a \ ('a,'b) tree \ ('a,'b) tree" +fixes inv :: "('a,'b) tree \ bool" +assumes inv_Leaf: "inv \\" +assumes set_join: "set_tree (join t1 x t2) = Set.insert x (set_tree t1 \ set_tree t2)" +assumes bst_join: + "\ bst l; bst r; \x \ set_tree l. x < k; \y \ set_tree r. k < y \ + \ bst (join l k r)" +assumes inv_join: "\ inv l; inv r \ \ inv (join l k r)" +assumes inv_Node: "\ inv (Node h l x r) \ \ inv l \ inv r" +begin + +declare set_join [simp] + +subsection "\split_min\" + +fun split_min :: "('a,'b) tree \ 'a \ ('a,'b) tree" where +"split_min (Node _ l x r) = + (if l = Leaf then (x,r) else let (m,l') = split_min l in (m, join l' x r))" + +lemma split_min_set: + "\ split_min t = (x,t'); t \ Leaf \ \ x \ set_tree t \ set_tree t = Set.insert x (set_tree t')" +proof(induction t arbitrary: t') + case Node thus ?case by(auto split: prod.splits if_splits dest: inv_Node) +next + case Leaf thus ?case by simp +qed + +lemma split_min_bst: + "\ split_min t = (x,t'); bst t; t \ Leaf \ \ bst t' \ (\x' \ set_tree t'. x < x')" +proof(induction t arbitrary: t') + case Node thus ?case by(fastforce simp: split_min_set bst_join split: prod.splits if_splits) +next + case Leaf thus ?case by simp +qed + +lemma split_min_inv: + "\ split_min t = (x,t'); inv t; t \ Leaf \ \ inv t'" +proof(induction t arbitrary: t') + case Node thus ?case by(auto simp: inv_join split: prod.splits if_splits dest: inv_Node) +next + case Leaf thus ?case by simp +qed + + +subsection "\join2\" + +definition join2 :: "('a,'b) tree \ ('a,'b) tree \ ('a,'b) tree" where +"join2 l r = (if r = Leaf then l else let (x,r') = split_min r in join l x r')" + +lemma set_join2[simp]: "set_tree (join2 l r) = set_tree l \ set_tree r" +by(simp add: join2_def split_min_set split: prod.split) + +lemma bst_join2: "\ bst l; bst r; \x \ set_tree l. \y \ set_tree r. x < y \ + \ bst (join2 l r)" +by(simp add: join2_def bst_join split_min_set split_min_bst split: prod.split) + +lemma inv_join2: "\ inv l; inv r \ \ inv (join2 l r)" +by(simp add: join2_def inv_join split_min_set split_min_inv split: prod.split) + + +subsection "\split\" + +fun split :: "('a,'b)tree \ 'a \ ('a,'b)tree \ bool \ ('a,'b)tree" where +"split Leaf k = (Leaf, False, Leaf)" | +"split (Node _ l a r) k = + (if k < a then let (l1,b,l2) = split l k in (l1, b, join l2 a r) else + if a < k then let (r1,b,r2) = split r k in (join l a r1, b, r2) + else (l, True, r))" + +lemma split: "split t k = (l,kin,r) \ bst t \ + set_tree l = {x \ set_tree t. x < k} \ set_tree r = {x \ set_tree t. k < x} + \ (kin = (k \ set_tree t)) \ bst l \ bst r" +proof(induction t arbitrary: l kin r) + case Leaf thus ?case by simp +next + case Node thus ?case by(force split!: prod.splits if_splits intro!: bst_join) +qed + +lemma split_inv: "split t k = (l,kin,r) \ inv t \ inv l \ inv r" +proof(induction t arbitrary: l kin r) + case Leaf thus ?case by simp +next + case Node + thus ?case by(force simp: inv_join split!: prod.splits if_splits dest!: inv_Node) +qed + +declare split.simps[simp del] + + +subsection "\insert\" + +definition insert :: "'a \ ('a,'b) tree \ ('a,'b) tree" where +"insert k t = (let (l,_,r) = split t k in join l k r)" + +lemma set_tree_insert: "bst t \ set_tree (insert x t) = Set.insert x (set_tree t)" +by(auto simp add: insert_def split split: prod.split) + +lemma bst_insert: "bst t \ bst (insert x t)" +by(auto simp add: insert_def bst_join dest: split split: prod.split) + +lemma inv_insert: "inv t \ inv (insert x t)" +by(force simp: insert_def inv_join dest: split_inv split: prod.split) + + +subsection "\delete\" + +definition delete :: "'a \ ('a,'b) tree \ ('a,'b) tree" where +"delete k t = (let (l,_,r) = split t k in join2 l r)" + +lemma set_tree_delete: "bst t \ set_tree (delete k t) = set_tree t - {k}" +by(auto simp: delete_def split split: prod.split) + +lemma bst_delete: "bst t \ bst (delete x t)" +by(force simp add: delete_def intro: bst_join2 dest: split split: prod.split) + +lemma inv_delete: "inv t \ inv (delete x t)" +by(force simp: delete_def inv_join2 dest: split_inv split: prod.split) + + +subsection "\union\" + +fun union :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +"union t1 t2 = + (if t1 = Leaf then t2 else + if t2 = Leaf then t1 else + case t1 of Node _ l1 k r1 \ + let (l2,_ ,r2) = split t2 k; + l' = union l1 l2; r' = union r1 r2 + in join l' k r')" + +declare union.simps [simp del] + +lemma set_tree_union: "bst t2 \ set_tree (union t1 t2) = set_tree t1 \ set_tree t2" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + then show ?case + by (auto simp: union.simps[of t1 t2] split split: tree.split prod.split) +qed + +lemma bst_union: "\ bst t1; bst t2 \ \ bst (union t1 t2)" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: union.simps[of t1 t2] set_tree_union split intro!: bst_join + split: tree.split prod.split) +qed + +lemma inv_union: "\ inv t1; inv t2 \ \ inv (union t1 t2)" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + thus ?case + by(auto simp:union.simps[of t1 t2] inv_join split_inv + split!: tree.split prod.split dest: inv_Node) +qed + +subsection "\inter\" + +fun inter :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +"inter t1 t2 = + (if t1 = Leaf then Leaf else + if t2 = Leaf then Leaf else + case t1 of Node _ l1 k r1 \ + let (l2,kin,r2) = split t2 k; + l' = inter l1 l2; r' = inter r1 r2 + in if kin then join l' k r' else join2 l' r')" + +declare inter.simps [simp del] + +lemma set_tree_inter: + "\ bst t1; bst t2 \ \ set_tree (inter t1 t2) = set_tree t1 \ set_tree t2" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + show ?case + proof (cases t1) + case Leaf thus ?thesis by (simp add: inter.simps) + next + case [simp]: (Node _ l1 k r1) + show ?thesis + proof (cases "t2 = Leaf") + case True thus ?thesis by (simp add: inter.simps) + next + case False + let ?L1 = "set_tree l1" let ?R1 = "set_tree r1" + have *: "k \ ?L1 \ ?R1" using \bst t1\ by (fastforce) + obtain l2 kin r2 where sp: "split t2 k = (l2,kin,r2)" using prod_cases3 by blast + let ?L2 = "set_tree l2" let ?R2 = "set_tree r2" let ?K = "if kin then {k} else {}" + have t2: "set_tree t2 = ?L2 \ ?R2 \ ?K" and + **: "?L2 \ ?R2 = {}" "k \ ?L2 \ ?R2" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" + using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force, force) + have IHl: "set_tree (inter l1 l2) = set_tree l1 \ set_tree l2" + using "1.IH"(1)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have IHr: "set_tree (inter r1 r2) = set_tree r1 \ set_tree r2" + using "1.IH"(2)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have "set_tree t1 \ set_tree t2 = (?L1 \ ?R1 \ {k}) \ (?L2 \ ?R2 \ ?K)" + by(simp add: t2) + also have "\ = (?L1 \ ?L2) \ (?R1 \ ?R2) \ ?K" + using * ** by auto + also have "\ = set_tree (inter t1 t2)" + using IHl IHr sp inter.simps[of t1 t2] False by(simp) + finally show ?thesis by simp + qed + qed +qed + +lemma bst_inter: "\ bst t1; bst t2 \ \ bst (inter t1 t2)" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: inter.simps[of t1 t2] set_tree_inter split Let_def + intro!: bst_join bst_join2 split: tree.split prod.split) +qed + +lemma inv_inter: "\ inv t1; inv t2 \ \ inv (inter t1 t2)" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + thus ?case + by(auto simp: inter.simps[of t1 t2] inv_join inv_join2 split_inv Let_def + split!: tree.split prod.split dest: inv_Node) +qed + +subsection "\diff\" + +fun diff :: "('a,'b)tree \ ('a,'b)tree \ ('a,'b)tree" where +"diff t1 t2 = + (if t1 = Leaf then Leaf else + if t2 = Leaf then t1 else + case t2 of Node _ l2 k r2 \ + let (l1,_,r1) = split t1 k; + l' = diff l1 l2; r' = diff r1 r2 + in join2 l' r')" + +declare diff.simps [simp del] + +lemma set_tree_diff: + "\ bst t1; bst t2 \ \ set_tree (diff t1 t2) = set_tree t1 - set_tree t2" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + show ?case + proof (cases t2) + case Leaf thus ?thesis by (simp add: diff.simps) + next + case [simp]: (Node _ l2 k r2) + show ?thesis + proof (cases "t1 = Leaf") + case True thus ?thesis by (simp add: diff.simps) + next + case False + let ?L2 = "set_tree l2" let ?R2 = "set_tree r2" + obtain l1 kin r1 where sp: "split t1 k = (l1,kin,r1)" using prod_cases3 by blast + let ?L1 = "set_tree l1" let ?R1 = "set_tree r1" let ?K = "if kin then {k} else {}" + have t1: "set_tree t1 = ?L1 \ ?R1 \ ?K" and + **: "k \ ?L1 \ ?R1" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" + using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force) + have IHl: "set_tree (diff l1 l2) = set_tree l1 - set_tree l2" + using "1.IH"(1)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have IHr: "set_tree (diff r1 r2) = set_tree r1 - set_tree r2" + using "1.IH"(2)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have "set_tree t1 - set_tree t2 = (?L1 \ ?R1) - (?L2 \ ?R2 \ {k})" + by(simp add: t1) + also have "\ = (?L1 - ?L2) \ (?R1 - ?R2)" + using ** by auto + also have "\ = set_tree (diff t1 t2)" + using IHl IHr sp diff.simps[of t1 t2] False by(simp) + finally show ?thesis by simp + qed + qed +qed + +lemma bst_diff: "\ bst t1; bst t2 \ \ bst (diff t1 t2)" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: diff.simps[of t1 t2] set_tree_diff split Let_def + intro!: bst_join bst_join2 split: tree.split prod.split) +qed + +lemma inv_diff: "\ inv t1; inv t2 \ \ inv (diff t1 t2)" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + thus ?case + by(auto simp: diff.simps[of t1 t2] inv_join inv_join2 split_inv Let_def + split!: tree.split prod.split dest: inv_Node) +qed + +text \Locale @{locale Set2_BST2_Join} implements locale @{locale Set2}:\ + +sublocale Set2 +where empty = Leaf and insert = insert and delete = delete and isin = isin +and union = union and inter = inter and diff = diff +and set = set_tree and invar = "\t. inv t \ bst t" +proof (standard, goal_cases) + case 1 show ?case by (simp) +next + case 2 thus ?case by(simp add: isin_iff) +next + case 3 thus ?case by (simp add: set_tree_insert) +next + case 4 thus ?case by (simp add: set_tree_delete) +next + case 5 thus ?case by (simp add: inv_Leaf) +next + case 6 thus ?case by (simp add: bst_insert inv_insert) +next + case 7 thus ?case by (simp add: bst_delete inv_delete) +next + case 8 thus ?case by(simp add: set_tree_union) +next + case 9 thus ?case by(simp add: set_tree_inter) +next + case 10 thus ?case by(simp add: set_tree_diff) +next + case 11 thus ?case by (simp add: bst_union inv_union) +next + case 12 thus ?case by (simp add: bst_inter inv_inter) +next + case 13 thus ?case by (simp add: bst_diff inv_diff) +qed + +end + +interpretation unbal: Set2_BST2_Join +where join = "\l x r. Node () l x r" and inv = "\t. True" +proof (standard, goal_cases) + case 1 show ?case by simp +next + case 2 thus ?case by simp +next + case 3 thus ?case by simp +next + case 4 thus ?case by simp +next + case 5 thus ?case by simp +qed + +end \ No newline at end of file diff -r aaa31cd0caef -r f13796496e82 src/HOL/Data_Structures/Set2_BST2_Join_RBT.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Set2_BST2_Join_RBT.thy Sun Apr 08 12:14:00 2018 +0200 @@ -0,0 +1,281 @@ +(* Author: Tobias Nipkow *) + +section "Join-Based BST2 Implementation of Sets via RBTs" + +theory Set2_BST2_Join_RBT +imports + Set2_BST2_Join + RBT_Set +begin + +subsection "Code" + +(* +WARNING baliL and baliR look symmetric but are not if you mirror them: +the first two clauses need to be swapped in one of the two. +Below we defined such a modified baliR. Is already modified in devel. +*) + +text \ +Function \joinL\ joins two trees (and an element). +Precondition: @{prop "bheight l \ bheight r"}. +Method: +Descend along the left spine of \r\ +until you find a subtree with the same \bheight\ as \l\, +then combine them into a new red node. +\ +fun joinL :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where +"joinL l x r = + (if bheight l = bheight r then R l x r + else case r of + B l' x' r' \ baliL (joinL l x l') x' r' | + R l' x' r' \ R (joinL l x l') x' r')" + +(* rm after isabelle release > 2017 because then in distribution *) +fun baliR :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where +"baliR t1 a1 (R t2 a2 (R t3 a3 t4)) = R (B t1 a1 t2) a2 (B t3 a3 t4)" | +"baliR t1 a1 (R (R t2 a2 t3) a3 t4) = R (B t1 a1 t2) a2 (B t3 a3 t4)" | +"baliR t1 a t2 = B t1 a t2" + +lemma inorder_baliR: + "inorder(baliR l a r) = inorder l @ a # inorder r" +by(cases "(l,a,r)" rule: baliR.cases) (auto) + +lemma invc_baliR: + "\invc l; invc2 r\ \ invc (baliR l a r)" +by (induct l a r rule: baliR.induct) auto + +lemma invh_baliR: + "\ invh l; invh r; bheight l = bheight r \ \ invh (baliR l a r)" +by (induct l a r rule: baliR.induct) auto +(* end rm *) + +fun joinR :: "'a rbt \ 'a \ 'a rbt \ 'a rbt" where +"joinR l x r = + (if bheight l \ bheight r then R l x r + else case l of + 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 +"join l x r = + (if bheight l > bheight r + then paint Black (joinR l x r) + else if bheight l < bheight r + then paint Black (joinL l x r) + else B l x r)" + +declare joinL.simps[simp del] +declare joinR.simps[simp del] + +text \ +One would expect @{const joinR} to be be completely dual to @{const joinL}. +Thus the condition should be @{prop"bheight l = bheight r"}. What we have done +is totalize the function. On the intended domain (@{prop "bheight l \ bheight r"}) +the two versions behave exactly the same, including complexity. Thus from a programmer's +perspective they are equivalent. However, not from a verifier's perspective: +the total version of @{const joinR} is easier +to reason about because lemmas about it may not require preconditions. In particular +@{prop"set_tree (joinR l x r) = Set.insert x (set_tree l \ set_tree r)"} +is provable outright and hence also +@{prop"set_tree (join l x r) = Set.insert x (set_tree l \ set_tree r)"}. +This is necessary because locale @{locale Set2_BST2_Join} unconditionally assumes +exactly that. Adding preconditions to this assumptions significantly complicates +the proofs within @{locale Set2_BST2_Join}, which we want to avoid. + +Why not work with the partial version of @{const joinR} and add the precondition +@{prop "bheight l \ bheight r"} to lemmas about @{const joinR}? After all, that is how +we worked with @{const joinL}, and @{const join} ensures that @{const joinL} and @{const joinR} +are only called under the respective precondition. But function @{const bheight} +makes the difference: it descends along the left spine, just like @{const joinL}. +Function @{const joinR}, however, descends along the right spine and thus @{const bheight} +may change all the time. Thus we would need the further precondition @{prop "invh l"}. +This is what we really wanted to avoid in order to satisfy the unconditional assumption +in @{locale Set2_BST2_Join}. +\ + +subsection "Properties" + +subsubsection "Color and height invariants" + +lemma invc2_joinL: + "\ invc l; invc r; bheight l \ bheight r \ \ + invc2 (joinL l x r) + \ (bheight l \ bheight r \ color r = Black \ invc(joinL l x r))" +proof (induct l x r rule: joinL.induct) + case (1 l x r) thus ?case + by(auto simp: invc_baliL invc2I joinL.simps[of l x r] split!: tree.splits if_splits) +qed + +lemma invc2_joinR: + "\ invc l; invh l; invc r; invh r; bheight l \ bheight r \ \ + invc2 (joinR l x r) + \ (bheight l \ bheight r \ color l = Black \ invc(joinR l x r))" +proof (induct l x r rule: joinR.induct) + case (1 l x r) thus ?case + by(fastforce simp: invc_baliR invc2I joinR.simps[of l x r] split!: tree.splits if_splits) +qed + +lemma bheight_joinL: + "\ invh l; invh r; bheight l \ bheight r \ \ bheight (joinL l x r) = bheight r" +proof (induct l x r rule: joinL.induct) + case (1 l x r) thus ?case + by(auto simp: bheight_baliL joinL.simps[of l x r] split!: tree.split) +qed + +lemma invh_joinL: + "\ invh l; invh r; bheight l \ bheight r \ \ invh (joinL l x r)" +proof (induct l x r rule: joinL.induct) + case (1 l x r) thus ?case + by(auto simp: invh_baliL bheight_joinL joinL.simps[of l x r] split!: tree.split color.split) +qed + +lemma bheight_baliR: + "bheight l = bheight r \ bheight (baliR l a r) = Suc (bheight l)" +by (cases "(l,a,r)" rule: baliR.cases) auto + +lemma bheight_joinR: + "\ invh l; invh r; bheight l \ bheight r \ \ bheight (joinR l x r) = bheight l" +proof (induct l x r rule: joinR.induct) + case (1 l x r) thus ?case + by(fastforce simp: bheight_baliR joinR.simps[of l x r] split!: tree.split) +qed + +lemma invh_joinR: + "\ invh l; invh r; bheight l \ bheight r \ \ invh (joinR l x r)" +proof (induct l x r rule: joinR.induct) + case (1 l x r) thus ?case + by(fastforce simp: invh_baliR bheight_joinR joinR.simps[of l x r] + split!: tree.split color.split) +qed + +(* 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) + +text \To make sure the the black height is not increased unnecessarily:\ + +lemma bheight_paint_Black: "bheight(paint Black t) \ bheight t + 1" +by(cases t) auto + +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) + + +subsubsection "Inorder properties" + +text "Currently unused. Instead @{const set_tree} and @{const bst} properties are proved directly." + +lemma inorder_joinL: "bheight l \ bheight r \ inorder(joinL l x r) = inorder l @ x # inorder r" +proof(induction l x r rule: joinL.induct) + case (1 l x r) + thus ?case by(auto simp: inorder_baliL joinL.simps[of l x r] split!: tree.splits color.splits) +qed + +lemma inorder_joinR: + "inorder(joinR l x r) = inorder l @ x # inorder r" +proof(induction l x r rule: joinR.induct) + case (1 l x r) + thus ?case by (force simp: inorder_baliR joinR.simps[of l x r] split!: tree.splits color.splits) +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 + dest!: arg_cong[where f = inorder]) + + +subsubsection "Set and bst properties" + +lemma set_baliL: + "set_tree(baliL l a r) = Set.insert a (set_tree l \ set_tree r)" +by(cases "(l,a,r)" rule: baliL.cases) (auto) + +lemma set_joinL: + "bheight l \ bheight r \ set_tree (joinL l x r) = Set.insert x (set_tree l \ set_tree r)" +proof(induction l x r rule: joinL.induct) + case (1 l x r) + thus ?case by(auto simp: set_baliL joinL.simps[of l x r] split!: tree.splits color.splits) +qed + +lemma set_baliR: + "set_tree(baliR l a r) = Set.insert a (set_tree l \ set_tree r)" +by(cases "(l,a,r)" rule: baliR.cases) (auto) + +lemma set_joinR: + "set_tree (joinR l x r) = Set.insert x (set_tree l \ set_tree r)" +proof(induction l x r rule: joinR.induct) + case (1 l x r) + thus ?case by(force simp: set_baliR joinR.simps[of l x r] split!: tree.splits color.splits) +qed + +lemma set_paint: "set_tree (paint c t) = set_tree t" +by (cases t) auto + +lemma set_join: "set_tree (join l x r) = Set.insert x (set_tree l \ set_tree r)" +by(simp add: set_joinL set_joinR set_paint) + +lemma bst_baliL: + "\bst l; bst r; \x\set_tree l. x < k; \x\set_tree r. k < x\ + \ bst (baliL l k r)" +by(cases "(l,k,r)" rule: baliL.cases) (auto simp: ball_Un) + +lemma bst_baliR: + "\bst l; bst r; \x\set_tree l. x < k; \x\set_tree r. k < x\ + \ bst (baliR l k r)" +by(cases "(l,k,r)" rule: baliR.cases) (auto simp: ball_Un) + +lemma bst_joinL: + "\bst l; bst r; \x\set_tree l. x < k; \y\set_tree r. k < y; bheight l \ bheight r\ + \ bst (joinL l k r)" +proof(induction l k r rule: joinL.induct) + case (1 l x r) + thus ?case + by(auto simp: set_baliL joinL.simps[of l x r] set_joinL ball_Un intro!: bst_baliL + split!: tree.splits color.splits) +qed + +lemma bst_joinR: + "\bst l; bst r; \x\set_tree l. x < k; \y\set_tree r. k < y \ + \ bst (joinR l k r)" +proof(induction l k r rule: joinR.induct) + case (1 l x r) + thus ?case + by(auto simp: set_baliR joinR.simps[of l x r] set_joinR ball_Un intro!: bst_baliR + split!: tree.splits color.splits) +qed + +lemma bst_paint: "bst (paint c t) = bst t" +by(cases t) auto + +lemma bst_join: + "\bst l; bst r; \x\set_tree l. x < k; \y\set_tree r. k < y \ + \ bst (join l k r)" +by(auto simp: bst_paint bst_joinL bst_joinR) + + +subsubsection "Interpretation of @{locale Set2_BST2_Join} with Red-Black Tree" + +global_interpretation RBT: Set2_BST2_Join +where join = join and inv = "\t. invc t \ invh t" +defines insert_rbt = RBT.insert and delete_rbt = RBT.delete and split_rbt = RBT.split +and join2_rbt = RBT.join2 and split_min_rbt = RBT.split_min +proof (standard, goal_cases) + case 1 show ?case by simp +next + case 2 show ?case by (rule set_join) +next + case 3 thus ?case by (rule bst_join) +next + case 4 thus ?case + by (simp add: invc2_joinL invc2_joinR invc_paint_Black invh_joinL invh_joinR invh_paint) +next + case 5 thus ?case by simp +qed + +text \The invariant does not guarantee that the root node is black. This is not required +to guarantee that the height is logarithmic in the size --- Exercise.\ + +end \ No newline at end of file diff -r aaa31cd0caef -r f13796496e82 src/HOL/Data_Structures/Set2_BST_Join.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Data_Structures/Set2_BST_Join.thy Sun Apr 08 12:14:00 2018 +0200 @@ -0,0 +1,364 @@ +(* Author: Tobias Nipkow *) + +section "Join-Based BST Implementation of Sets" + +theory Set2_BST_Join +imports + "HOL-Library.Tree" + Cmp + Set_Specs +begin + +text \This theory implements the set operations \insert\, \delete\, +\union\, \inter\section and \diff\erence. The implementation is based on binary search trees. +All operations are reduced to a single operation \join l x r\ that joins two BSTs \l\ and \r\ +and an element \x\ such that \l < x < r\. + +This theory illustrates the idea but is not suitable for an efficient implementation where +\join\ balances the tree in some manner because type @{typ "'a tree"} in theory @{theory Tree} +has no additional fields for recording balance information. See theory \Set2_BST2_Join\ for that.\ + +text \Function \isin\ can also be expressed via \join\ but this is more direct:\ + +fun isin :: "('a::linorder) tree \ 'a \ bool" where +"isin Leaf x = False" | +"isin (Node l a r) x = + (case cmp x a of + LT \ isin l x | + EQ \ True | + GT \ isin r x)" + +lemma isin_set: "bst t \ isin t x = (x \ set_tree t)" +by (induction t) (auto) + + +locale Set2_BST_Join = +fixes join :: "('a::linorder) tree \ 'a \ 'a tree \ 'a tree" +assumes set_join: "set_tree (join t1 x t2) = Set.insert x (set_tree t1 \ set_tree t2)" +assumes bst_join: + "\ bst l; bst r; \x \ set_tree l. x < k; \y \ set_tree r. k < y \ + \ bst (join l k r)" +fixes inv :: "'a tree \ bool" +assumes inv_Leaf: "inv \\" +assumes inv_join: "\ inv l; inv r \ \ inv (join l k r)" +assumes inv_Node: "\ inv (Node l x r) \ \ inv l \ inv r" +begin + +declare set_join [simp] + +subsection "\split_min\" + +fun split_min :: "'a tree \ 'a \ 'a tree" where +"split_min (Node l x r) = + (if l = Leaf then (x,r) else let (m,l') = split_min l in (m, join l' x r))" + +lemma split_min_set: + "\ split_min t = (x,t'); t \ Leaf \ \ + x \ set_tree t \ set_tree t = Set.insert x (set_tree t')" +proof(induction t arbitrary: t') + case Node thus ?case by(auto split: prod.splits if_splits) +next + case Leaf thus ?case by simp +qed + +lemma split_min_bst: + "\ split_min t = (x,t'); bst t; t \ Leaf \ \ bst t' \ (\x' \ set_tree t'. x < x')" +proof(induction t arbitrary: t') + case Node thus ?case by(fastforce simp: split_min_set bst_join split: prod.splits if_splits) +next + case Leaf thus ?case by simp +qed + +lemma split_min_inv: + "\ split_min t = (x,t'); inv t; t \ Leaf \ \ inv t'" +proof(induction t arbitrary: t') + case Node thus ?case by(auto simp: inv_join split: prod.splits if_splits dest: inv_Node) +next + case Leaf thus ?case by simp +qed + + +subsection "\join2\" + +definition join2 :: "'a tree \ 'a tree \ 'a tree" where +"join2 l r = (if r = Leaf then l else let (x,r') = split_min r in join l x r')" + +lemma set_join2[simp]: "set_tree (join2 l r) = set_tree l \ set_tree r" +by(simp add: join2_def split_min_set split: prod.split) + +lemma bst_join2: "\ bst l; bst r; \x \ set_tree l. \y \ set_tree r. x < y \ + \ bst (join2 l r)" +by(simp add: join2_def bst_join split_min_set split_min_bst split: prod.split) + +lemma inv_join2: "\ inv l; inv r \ \ inv (join2 l r)" +by(simp add: join2_def inv_join split_min_set split_min_inv split: prod.split) + + +subsection "\split\" + +fun split :: "'a tree \ 'a \ 'a tree \ bool \ 'a tree" where +"split Leaf k = (Leaf, False, Leaf)" | +"split (Node l a r) k = + (case cmp k a of + LT \ let (l1,b,l2) = split l k in (l1, b, join l2 a r) | + GT \ let (r1,b,r2) = split r k in (join l a r1, b, r2) | + EQ \ (l, True, r))" + +lemma split: "split t k = (l,kin,r) \ bst t \ + set_tree l = {x \ set_tree t. x < k} \ set_tree r = {x \ set_tree t. k < x} + \ (kin = (k \ set_tree t)) \ bst l \ bst r" +proof(induction t arbitrary: l kin r) + case Leaf thus ?case by simp +next + case Node thus ?case by(force split!: prod.splits if_splits intro!: bst_join) +qed + +lemma split_inv: "split t k = (l,kin,r) \ inv t \ inv l \ inv r" +proof(induction t arbitrary: l kin r) + case Leaf thus ?case by simp +next + case Node + thus ?case by(force simp: inv_join split!: prod.splits if_splits dest!: inv_Node) +qed + +declare split.simps[simp del] + + +subsection "\insert\" + +definition insert :: "'a \ 'a tree \ 'a tree" where +"insert k t = (let (l,_,r) = split t k in join l k r)" + +lemma set_tree_insert: "bst t \ set_tree (insert x t) = Set.insert x (set_tree t)" +by(auto simp add: insert_def split split: prod.split) + +lemma bst_insert: "bst t \ bst (insert x t)" +by(auto simp add: insert_def bst_join dest: split split: prod.split) + +lemma inv_insert: "inv t \ inv (insert x t)" +by(force simp: insert_def inv_join dest: split_inv split: prod.split) + + +subsection "\delete\" + +definition delete :: "'a \ 'a tree \ 'a tree" where +"delete k t = (let (l,_,r) = split t k in join2 l r)" + +lemma set_tree_delete: "bst t \ set_tree (delete k t) = set_tree t - {k}" +by(auto simp: delete_def split split: prod.split) + +lemma bst_delete: "bst t \ bst (delete x t)" +by(force simp add: delete_def intro: bst_join2 dest: split split: prod.split) + +lemma inv_delete: "inv t \ inv (delete x t)" +by(force simp: delete_def inv_join2 dest: split_inv split: prod.split) + + +subsection "\union\" + +fun union :: "'a tree \ 'a tree \ 'a tree" where +"union t1 t2 = + (if t1 = Leaf then t2 else + if t2 = Leaf then t1 else + case t1 of Node l1 k r1 \ + let (l2,_ ,r2) = split t2 k; + l' = union l1 l2; r' = union r1 r2 + in join l' k r')" + +declare union.simps [simp del] + +lemma set_tree_union: "bst t2 \ set_tree (union t1 t2) = set_tree t1 \ set_tree t2" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + then show ?case + by (auto simp: union.simps[of t1 t2] split split: tree.split prod.split) +qed + +lemma bst_union: "\ bst t1; bst t2 \ \ bst (union t1 t2)" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: union.simps[of t1 t2] set_tree_union split intro!: bst_join + split: tree.split prod.split) +qed + +lemma inv_union: "\ inv t1; inv t2 \ \ inv (union t1 t2)" +proof(induction t1 t2 rule: union.induct) + case (1 t1 t2) + thus ?case + by(auto simp:union.simps[of t1 t2] inv_join split_inv + split!: tree.split prod.split dest: inv_Node) +qed + +subsection "\inter\" + +fun inter :: "'a tree \ 'a tree \ 'a tree" where +"inter t1 t2 = + (if t1 = Leaf then Leaf else + if t2 = Leaf then Leaf else + case t1 of Node l1 k r1 \ + let (l2,kin,r2) = split t2 k; + l' = inter l1 l2; r' = inter r1 r2 + in if kin then join l' k r' else join2 l' r')" + +declare inter.simps [simp del] + +lemma set_tree_inter: + "\ bst t1; bst t2 \ \ set_tree (inter t1 t2) = set_tree t1 \ set_tree t2" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + show ?case + proof (cases t1) + case Leaf thus ?thesis by (simp add: inter.simps) + next + case [simp]: (Node l1 k r1) + show ?thesis + proof (cases "t2 = Leaf") + case True thus ?thesis by (simp add: inter.simps) + next + case False + let ?L1 = "set_tree l1" let ?R1 = "set_tree r1" + have *: "k \ ?L1 \ ?R1" using \bst t1\ by (fastforce) + obtain l2 kin r2 where sp: "split t2 k = (l2,kin,r2)" using prod_cases3 by blast + let ?L2 = "set_tree l2" let ?R2 = "set_tree r2" let ?K = "if kin then {k} else {}" + have t2: "set_tree t2 = ?L2 \ ?R2 \ ?K" and + **: "?L2 \ ?R2 = {}" "k \ ?L2 \ ?R2" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" + using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force, force) + have IHl: "set_tree (inter l1 l2) = set_tree l1 \ set_tree l2" + using "1.IH"(1)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have IHr: "set_tree (inter r1 r2) = set_tree r1 \ set_tree r2" + using "1.IH"(2)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have "set_tree t1 \ set_tree t2 = (?L1 \ ?R1 \ {k}) \ (?L2 \ ?R2 \ ?K)" + by(simp add: t2) + also have "\ = (?L1 \ ?L2) \ (?R1 \ ?R2) \ ?K" + using * ** by auto + also have "\ = set_tree (inter t1 t2)" + using IHl IHr sp inter.simps[of t1 t2] False by(simp) + finally show ?thesis by simp + qed + qed +qed + +lemma bst_inter: "\ bst t1; bst t2 \ \ bst (inter t1 t2)" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: inter.simps[of t1 t2] set_tree_inter split Let_def + intro!: bst_join bst_join2 split: tree.split prod.split) +qed + +lemma inv_inter: "\ inv t1; inv t2 \ \ inv (inter t1 t2)" +proof(induction t1 t2 rule: inter.induct) + case (1 t1 t2) + thus ?case + by(auto simp: inter.simps[of t1 t2] inv_join inv_join2 split_inv Let_def + split!: tree.split prod.split dest: inv_Node) +qed + +subsection "\diff\" + +fun diff :: "'a tree \ 'a tree \ 'a tree" where +"diff t1 t2 = + (if t1 = Leaf then Leaf else + if t2 = Leaf then t1 else + case t2 of Node l2 k r2 \ + let (l1,_,r1) = split t1 k; + l' = diff l1 l2; r' = diff r1 r2 + in join2 l' r')" + +declare diff.simps [simp del] + +lemma set_tree_diff: + "\ bst t1; bst t2 \ \ set_tree (diff t1 t2) = set_tree t1 - set_tree t2" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + show ?case + proof (cases t2) + case Leaf thus ?thesis by (simp add: diff.simps) + next + case [simp]: (Node l2 k r2) + show ?thesis + proof (cases "t1 = Leaf") + case True thus ?thesis by (simp add: diff.simps) + next + case False + let ?L2 = "set_tree l2" let ?R2 = "set_tree r2" + obtain l1 kin r1 where sp: "split t1 k = (l1,kin,r1)" using prod_cases3 by blast + let ?L1 = "set_tree l1" let ?R1 = "set_tree r1" let ?K = "if kin then {k} else {}" + have t1: "set_tree t1 = ?L1 \ ?R1 \ ?K" and + **: "k \ ?L1 \ ?R1" "?L1 \ ?R2 = {}" "?L2 \ ?R1 = {}" + using split[OF sp] \bst t1\ \bst t2\ by (force, force, force, force) + have IHl: "set_tree (diff l1 l2) = set_tree l1 - set_tree l2" + using "1.IH"(1)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have IHr: "set_tree (diff r1 r2) = set_tree r1 - set_tree r2" + using "1.IH"(2)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp + have "set_tree t1 - set_tree t2 = (?L1 \ ?R1) - (?L2 \ ?R2 \ {k})" + by(simp add: t1) + also have "\ = (?L1 - ?L2) \ (?R1 - ?R2)" + using ** by auto + also have "\ = set_tree (diff t1 t2)" + using IHl IHr sp diff.simps[of t1 t2] False by(simp) + finally show ?thesis by simp + qed + qed +qed + +lemma bst_diff: "\ bst t1; bst t2 \ \ bst (diff t1 t2)" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + thus ?case + by(fastforce simp: diff.simps[of t1 t2] set_tree_diff split Let_def + intro!: bst_join bst_join2 split: tree.split prod.split) +qed + +lemma inv_diff: "\ inv t1; inv t2 \ \ inv (diff t1 t2)" +proof(induction t1 t2 rule: diff.induct) + case (1 t1 t2) + thus ?case + by(auto simp: diff.simps[of t1 t2] inv_join inv_join2 split_inv Let_def + split!: tree.split prod.split dest: inv_Node) +qed + +text \Locale @{locale Set2_BST_Join} implements locale @{locale Set2}:\ + +sublocale Set2 +where empty = Leaf and isin = isin and insert = insert and delete = delete +and union = union and inter = inter and diff = diff +and set = set_tree and invar = "\t. bst t \ inv t" +proof (standard, goal_cases) + case 1 show ?case by simp +next + case 2 thus ?case by (simp add: isin_set) +next + case 3 thus ?case by (simp add: set_tree_insert) +next + case 4 thus ?case by (simp add: set_tree_delete) +next + case 5 thus ?case by (simp add: inv_Leaf) +next + case 6 thus ?case by (simp add: inv_insert bst_insert) +next + case 7 thus ?case by (simp add: inv_delete bst_delete) +next + case 8 thus ?case by (simp add: set_tree_union) +next + case 9 thus ?case by (simp add: set_tree_inter) +next + case 10 thus ?case by (simp add: set_tree_diff) +next + case 11 thus ?case by (simp add: bst_union inv_union) +next + case 12 thus ?case by (simp add: bst_inter inv_inter) +next + case 13 thus ?case by (simp add: bst_diff inv_diff) +qed + +end (* Set2_BST_Join *) + +text \Interpretation of @{locale Set2_BST_Join} with unbalanced binary trees:\ + +interpretation Set2_BST_Join where join = Node and inv = "\t. True" +proof (standard, goal_cases) +qed auto + +end \ No newline at end of file diff -r aaa31cd0caef -r f13796496e82 src/HOL/Data_Structures/document/root.bib --- a/src/HOL/Data_Structures/document/root.bib Sun Apr 08 11:05:52 2018 +0200 +++ b/src/HOL/Data_Structures/document/root.bib Sun Apr 08 12:14:00 2018 +0200 @@ -1,8 +1,33 @@ +@techreport{Adams-TR92,author="Stephen Adams", +title="Implementing Sets Efficiently in a Functional Language", +institution="University of Southampton, Department of Electronics and Computer Science",number="CSTR 92-10",year=1992} + +@article{Adams-JFP93, + author = {Stephen Adams}, + title = {Efficient Sets - {A} Balancing Act}, + journal = {J. Funct. Program.}, + volume = {3}, + number = {4}, + pages = {553--561}, + year = {1993} +} + @inproceedings{Andersson-WADS93,author={Arne Andersson}, title={Balanced search trees made simple},pages={60--71},year=1993, booktitle={Algorithms and Data Structures (WADS '93)}, series={LNCS},volume={709},publisher={Springer}} +@inproceedings{BlellochFS-SPAA16, + author = {Guy E. Blelloch and + Daniel Ferizovic and + Yihan Sun}, + title = {Just Join for Parallel Ordered Sets}, + booktitle = {{SPAA}}, + pages = {253--264}, + publisher = {{ACM}}, + year = {2016} +} + @phdthesis{Crane72,author={Clark A. Crane}, title={Linear Lists and Prorty Queues as Balanced Binary Trees}, school={Computer Science Department, Stanford University},year=1972} diff -r aaa31cd0caef -r f13796496e82 src/HOL/Data_Structures/document/root.tex --- a/src/HOL/Data_Structures/document/root.tex Sun Apr 08 11:05:52 2018 +0200 +++ b/src/HOL/Data_Structures/document/root.tex Sun Apr 08 12:14:00 2018 +0200 @@ -65,6 +65,10 @@ They were invented by Sleator and Tarjan \cite{SleatorT-JACM85}. Our formalisation follows Schoenmakers \cite{Schoenmakers-IPL93}. +\paragraph{Join-based BSTs} +They were invented by Adams \cite{Adams-TR92,Adams-JFP93} +and analyzed by Blelloch \emph{et al.} \cite{BlellochFS-SPAA16}. + \paragraph{Leftist heaps} They were invented by Crane \cite{Crane72}. A first functional implementation is due to N\'u\~{n}ez \emph{et al.}~\cite{NunezPP95}.