src/HOL/Data_Structures/Set2_BST2_Join.thy
changeset 68261 035c78bb0a66
parent 68260 61188c781cdd
child 68262 d231238bd977
equal deleted inserted replaced
68260:61188c781cdd 68261:035c78bb0a66
     1 (* Author: Tobias Nipkow *)
       
     2 
       
     3 section "Join-Based BST2 Implementation of Sets"
       
     4 
       
     5 theory Set2_BST2_Join
       
     6 imports
       
     7   Isin2
       
     8 begin
       
     9 
       
    10 text \<open>Based on theory @{theory Tree2}, otherwise same as theory @{file "Set2_BST_Join.thy"}.\<close>
       
    11 
       
    12 text \<open>This theory implements the set operations \<open>insert\<close>, \<open>delete\<close>,
       
    13 \<open>union\<close>, \<open>inter\<close>section and \<open>diff\<close>erence. The implementation is based on binary search trees.
       
    14 All operations are reduced to a single operation \<open>join l x r\<close> that joins two BSTs \<open>l\<close> and \<open>r\<close>
       
    15 and an element \<open>x\<close> such that \<open>l < x < r\<close>.
       
    16 
       
    17 The theory is based on theory @{theory Tree2} where nodes have an additional field.
       
    18 This field is ignored here but it means that this theory can be instantiated
       
    19 with red-black trees (see theory @{file "Set2_BST2_Join_RBT.thy"}) and other balanced trees.
       
    20 This approach is very concrete and fixes the type of trees.
       
    21 Alternatively, one could assume some abstract type @{typ 't} of trees with suitable decomposition
       
    22 and recursion operators on it.\<close>
       
    23 
       
    24 locale Set2_BST2_Join =
       
    25 fixes join :: "('a::linorder,'b) tree \<Rightarrow> 'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree"
       
    26 fixes inv :: "('a,'b) tree \<Rightarrow> bool"
       
    27 assumes inv_Leaf: "inv \<langle>\<rangle>"
       
    28 assumes set_join: "set_tree (join t1 x t2) = Set.insert x (set_tree t1 \<union> set_tree t2)"
       
    29 assumes bst_join:
       
    30   "\<lbrakk> bst l; bst r; \<forall>x \<in> set_tree l. x < k; \<forall>y \<in> set_tree r. k < y \<rbrakk>
       
    31   \<Longrightarrow> bst (join l k r)"
       
    32 assumes inv_join: "\<lbrakk> inv l; inv r \<rbrakk> \<Longrightarrow> inv (join l k r)"
       
    33 assumes inv_Node: "\<lbrakk> inv (Node h l x r) \<rbrakk> \<Longrightarrow> inv l \<and> inv r"
       
    34 begin
       
    35 
       
    36 declare set_join [simp]
       
    37 
       
    38 subsection "\<open>split_min\<close>"
       
    39 
       
    40 fun split_min :: "('a,'b) tree \<Rightarrow> 'a \<times> ('a,'b) tree" where
       
    41 "split_min (Node _ l x r) =
       
    42   (if l = Leaf then (x,r) else let (m,l') = split_min l in (m, join l' x r))"
       
    43 
       
    44 lemma split_min_set:
       
    45   "\<lbrakk> split_min t = (x,t');  t \<noteq> Leaf \<rbrakk> \<Longrightarrow> x \<in> set_tree t \<and> set_tree t = Set.insert x (set_tree t')"
       
    46 proof(induction t arbitrary: t')
       
    47   case Node thus ?case by(auto split: prod.splits if_splits dest: inv_Node)
       
    48 next
       
    49   case Leaf thus ?case by simp
       
    50 qed
       
    51 
       
    52 lemma split_min_bst:
       
    53   "\<lbrakk> split_min t = (x,t');  bst t;  t \<noteq> Leaf \<rbrakk> \<Longrightarrow>  bst t' \<and> (\<forall>x' \<in> set_tree t'. x < x')"
       
    54 proof(induction t arbitrary: t')
       
    55   case Node thus ?case by(fastforce simp: split_min_set bst_join split: prod.splits if_splits)
       
    56 next
       
    57   case Leaf thus ?case by simp
       
    58 qed
       
    59 
       
    60 lemma split_min_inv:
       
    61   "\<lbrakk> split_min t = (x,t');  inv t;  t \<noteq> Leaf \<rbrakk> \<Longrightarrow>  inv t'"
       
    62 proof(induction t arbitrary: t')
       
    63   case Node thus ?case by(auto simp: inv_join split: prod.splits if_splits dest: inv_Node)
       
    64 next
       
    65   case Leaf thus ?case by simp
       
    66 qed
       
    67 
       
    68 
       
    69 subsection "\<open>join2\<close>"
       
    70 
       
    71 definition join2 :: "('a,'b) tree \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
       
    72 "join2 l r = (if r = Leaf then l else let (x,r') = split_min r in join l x r')"
       
    73 
       
    74 lemma set_join2[simp]: "set_tree (join2 l r) = set_tree l \<union> set_tree r"
       
    75 by(simp add: join2_def split_min_set split: prod.split)
       
    76 
       
    77 lemma bst_join2: "\<lbrakk> bst l; bst r; \<forall>x \<in> set_tree l. \<forall>y \<in> set_tree r. x < y \<rbrakk>
       
    78   \<Longrightarrow> bst (join2 l r)"
       
    79 by(simp add: join2_def bst_join split_min_set split_min_bst split: prod.split)
       
    80 
       
    81 lemma inv_join2: "\<lbrakk> inv l; inv r \<rbrakk> \<Longrightarrow> inv (join2 l r)"
       
    82 by(simp add: join2_def inv_join split_min_set split_min_inv split: prod.split)
       
    83 
       
    84 
       
    85 subsection "\<open>split\<close>"
       
    86 
       
    87 fun split :: "('a,'b)tree \<Rightarrow> 'a \<Rightarrow> ('a,'b)tree \<times> bool \<times> ('a,'b)tree" where
       
    88 "split Leaf k = (Leaf, False, Leaf)" |
       
    89 "split (Node _ l a r) k =
       
    90   (if k < a then let (l1,b,l2) = split l k in (l1, b, join l2 a r) else
       
    91    if a < k then let (r1,b,r2) = split r k in (join l a r1, b, r2)
       
    92    else (l, True, r))"
       
    93 
       
    94 lemma split: "split t k = (l,kin,r) \<Longrightarrow> bst t \<Longrightarrow>
       
    95   set_tree l = {x \<in> set_tree t. x < k} \<and> set_tree r = {x \<in> set_tree t. k < x}
       
    96   \<and> (kin = (k \<in> set_tree t)) \<and> bst l \<and> bst r"
       
    97 proof(induction t arbitrary: l kin r)
       
    98   case Leaf thus ?case by simp
       
    99 next
       
   100   case Node thus ?case by(force split!: prod.splits if_splits intro!: bst_join)
       
   101 qed
       
   102 
       
   103 lemma split_inv: "split t k = (l,kin,r) \<Longrightarrow> inv t \<Longrightarrow> inv l \<and> inv r"
       
   104 proof(induction t arbitrary: l kin r)
       
   105   case Leaf thus ?case by simp
       
   106 next
       
   107   case Node
       
   108   thus ?case by(force simp: inv_join split!: prod.splits if_splits dest!: inv_Node)
       
   109 qed
       
   110 
       
   111 declare split.simps[simp del]
       
   112 
       
   113 
       
   114 subsection "\<open>insert\<close>"
       
   115 
       
   116 definition insert :: "'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
       
   117 "insert k t = (let (l,_,r) = split t k in join l k r)"
       
   118 
       
   119 lemma set_tree_insert: "bst t \<Longrightarrow> set_tree (insert x t) = Set.insert x (set_tree t)"
       
   120 by(auto simp add: insert_def split split: prod.split)
       
   121 
       
   122 lemma bst_insert: "bst t \<Longrightarrow> bst (insert x t)"
       
   123 by(auto simp add: insert_def bst_join dest: split split: prod.split)
       
   124 
       
   125 lemma inv_insert: "inv t \<Longrightarrow> inv (insert x t)"
       
   126 by(force simp: insert_def inv_join dest: split_inv split: prod.split)
       
   127 
       
   128 
       
   129 subsection "\<open>delete\<close>"
       
   130 
       
   131 definition delete :: "'a \<Rightarrow> ('a,'b) tree \<Rightarrow> ('a,'b) tree" where
       
   132 "delete k t = (let (l,_,r) = split t k in join2 l r)"
       
   133 
       
   134 lemma set_tree_delete: "bst t \<Longrightarrow> set_tree (delete k t) = set_tree t - {k}"
       
   135 by(auto simp: delete_def split split: prod.split)
       
   136 
       
   137 lemma bst_delete: "bst t \<Longrightarrow> bst (delete x t)"
       
   138 by(force simp add: delete_def intro: bst_join2 dest: split split: prod.split)
       
   139 
       
   140 lemma inv_delete: "inv t \<Longrightarrow> inv (delete x t)"
       
   141 by(force simp: delete_def inv_join2 dest: split_inv split: prod.split)
       
   142 
       
   143 
       
   144 subsection "\<open>union\<close>"
       
   145 
       
   146 fun union :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
       
   147 "union t1 t2 =
       
   148   (if t1 = Leaf then t2 else
       
   149    if t2 = Leaf then t1 else
       
   150    case t1 of Node _ l1 k r1 \<Rightarrow>
       
   151    let (l2,_ ,r2) = split t2 k;
       
   152        l' = union l1 l2; r' = union r1 r2
       
   153    in join l' k r')"
       
   154 
       
   155 declare union.simps [simp del]
       
   156 
       
   157 lemma set_tree_union: "bst t2 \<Longrightarrow> set_tree (union t1 t2) = set_tree t1 \<union> set_tree t2"
       
   158 proof(induction t1 t2 rule: union.induct)
       
   159   case (1 t1 t2)
       
   160   then show ?case
       
   161     by (auto simp: union.simps[of t1 t2] split split: tree.split prod.split)
       
   162 qed
       
   163 
       
   164 lemma bst_union: "\<lbrakk> bst t1; bst t2 \<rbrakk> \<Longrightarrow> bst (union t1 t2)"
       
   165 proof(induction t1 t2 rule: union.induct)
       
   166   case (1 t1 t2)
       
   167   thus ?case
       
   168     by(fastforce simp: union.simps[of t1 t2] set_tree_union split intro!: bst_join 
       
   169         split: tree.split prod.split)
       
   170 qed
       
   171 
       
   172 lemma inv_union: "\<lbrakk> inv t1; inv t2 \<rbrakk> \<Longrightarrow> inv (union t1 t2)"
       
   173 proof(induction t1 t2 rule: union.induct)
       
   174   case (1 t1 t2)
       
   175   thus ?case
       
   176     by(auto simp:union.simps[of t1 t2] inv_join split_inv
       
   177         split!: tree.split prod.split dest: inv_Node)
       
   178 qed
       
   179 
       
   180 subsection "\<open>inter\<close>"
       
   181 
       
   182 fun inter :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
       
   183 "inter t1 t2 =
       
   184   (if t1 = Leaf then Leaf else
       
   185    if t2 = Leaf then Leaf else
       
   186    case t1 of Node _ l1 k r1 \<Rightarrow>
       
   187    let (l2,kin,r2) = split t2 k;
       
   188        l' = inter l1 l2; r' = inter r1 r2
       
   189    in if kin then join l' k r' else join2 l' r')"
       
   190 
       
   191 declare inter.simps [simp del]
       
   192 
       
   193 lemma set_tree_inter:
       
   194   "\<lbrakk> bst t1; bst t2 \<rbrakk> \<Longrightarrow> set_tree (inter t1 t2) = set_tree t1 \<inter> set_tree t2"
       
   195 proof(induction t1 t2 rule: inter.induct)
       
   196   case (1 t1 t2)
       
   197   show ?case
       
   198   proof (cases t1)
       
   199     case Leaf thus ?thesis by (simp add: inter.simps)
       
   200   next
       
   201     case [simp]: (Node _ l1 k r1)
       
   202     show ?thesis
       
   203     proof (cases "t2 = Leaf")
       
   204       case True thus ?thesis by (simp add: inter.simps)
       
   205     next
       
   206       case False
       
   207       let ?L1 = "set_tree l1" let ?R1 = "set_tree r1"
       
   208       have *: "k \<notin> ?L1 \<union> ?R1" using \<open>bst t1\<close> by (fastforce)
       
   209       obtain l2 kin r2 where sp: "split t2 k = (l2,kin,r2)" using prod_cases3 by blast
       
   210       let ?L2 = "set_tree l2" let ?R2 = "set_tree r2" let ?K = "if kin then {k} else {}"
       
   211       have t2: "set_tree t2 = ?L2 \<union> ?R2 \<union> ?K" and
       
   212            **: "?L2 \<inter> ?R2 = {}" "k \<notin> ?L2 \<union> ?R2" "?L1 \<inter> ?R2 = {}" "?L2 \<inter> ?R1 = {}"
       
   213         using split[OF sp] \<open>bst t1\<close> \<open>bst t2\<close> by (force, force, force, force, force)
       
   214       have IHl: "set_tree (inter l1 l2) = set_tree l1 \<inter> set_tree l2"
       
   215         using "1.IH"(1)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp
       
   216       have IHr: "set_tree (inter r1 r2) = set_tree r1 \<inter> set_tree r2"
       
   217         using "1.IH"(2)[OF _ False _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp
       
   218       have "set_tree t1 \<inter> set_tree t2 = (?L1 \<union> ?R1 \<union> {k}) \<inter> (?L2 \<union> ?R2 \<union> ?K)"
       
   219         by(simp add: t2)
       
   220       also have "\<dots> = (?L1 \<inter> ?L2) \<union> (?R1 \<inter> ?R2) \<union> ?K"
       
   221         using * ** by auto
       
   222       also have "\<dots> = set_tree (inter t1 t2)"
       
   223       using IHl IHr sp inter.simps[of t1 t2] False by(simp)
       
   224       finally show ?thesis by simp
       
   225     qed
       
   226   qed
       
   227 qed
       
   228 
       
   229 lemma bst_inter: "\<lbrakk> bst t1; bst t2 \<rbrakk> \<Longrightarrow> bst (inter t1 t2)"
       
   230 proof(induction t1 t2 rule: inter.induct)
       
   231   case (1 t1 t2)
       
   232   thus ?case
       
   233     by(fastforce simp: inter.simps[of t1 t2] set_tree_inter split Let_def
       
   234         intro!: bst_join bst_join2 split: tree.split prod.split)
       
   235 qed
       
   236 
       
   237 lemma inv_inter: "\<lbrakk> inv t1; inv t2 \<rbrakk> \<Longrightarrow> inv (inter t1 t2)"
       
   238 proof(induction t1 t2 rule: inter.induct)
       
   239   case (1 t1 t2)
       
   240   thus ?case
       
   241     by(auto simp: inter.simps[of t1 t2] inv_join inv_join2 split_inv Let_def
       
   242         split!: tree.split prod.split dest: inv_Node)
       
   243 qed
       
   244 
       
   245 subsection "\<open>diff\<close>"
       
   246 
       
   247 fun diff :: "('a,'b)tree \<Rightarrow> ('a,'b)tree \<Rightarrow> ('a,'b)tree" where
       
   248 "diff t1 t2 =
       
   249   (if t1 = Leaf then Leaf else
       
   250    if t2 = Leaf then t1 else
       
   251    case t2 of Node _ l2 k r2 \<Rightarrow>
       
   252    let (l1,_,r1) = split t1 k;
       
   253        l' = diff l1 l2; r' = diff r1 r2
       
   254    in join2 l' r')"
       
   255 
       
   256 declare diff.simps [simp del]
       
   257 
       
   258 lemma set_tree_diff:
       
   259   "\<lbrakk> bst t1; bst t2 \<rbrakk> \<Longrightarrow> set_tree (diff t1 t2) = set_tree t1 - set_tree t2"
       
   260 proof(induction t1 t2 rule: diff.induct)
       
   261   case (1 t1 t2)
       
   262   show ?case
       
   263   proof (cases t2)
       
   264     case Leaf thus ?thesis by (simp add: diff.simps)
       
   265   next
       
   266     case [simp]: (Node _ l2 k r2)
       
   267     show ?thesis
       
   268     proof (cases "t1 = Leaf")
       
   269       case True thus ?thesis by (simp add: diff.simps)
       
   270     next
       
   271       case False
       
   272       let ?L2 = "set_tree l2" let ?R2 = "set_tree r2"
       
   273       obtain l1 kin r1 where sp: "split t1 k = (l1,kin,r1)" using prod_cases3 by blast
       
   274       let ?L1 = "set_tree l1" let ?R1 = "set_tree r1" let ?K = "if kin then {k} else {}"
       
   275       have t1: "set_tree t1 = ?L1 \<union> ?R1 \<union> ?K" and
       
   276            **: "k \<notin> ?L1 \<union> ?R1" "?L1 \<inter> ?R2 = {}" "?L2 \<inter> ?R1 = {}"
       
   277         using split[OF sp] \<open>bst t1\<close> \<open>bst t2\<close> by (force, force, force, force)
       
   278       have IHl: "set_tree (diff l1 l2) = set_tree l1 - set_tree l2"
       
   279         using "1.IH"(1)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp
       
   280       have IHr: "set_tree (diff r1 r2) = set_tree r1 - set_tree r2"
       
   281         using "1.IH"(2)[OF False _ _ sp[symmetric]] "1.prems"(1,2) split[OF sp] by simp
       
   282       have "set_tree t1 - set_tree t2 = (?L1 \<union> ?R1) - (?L2 \<union> ?R2  \<union> {k})"
       
   283         by(simp add: t1)
       
   284       also have "\<dots> = (?L1 - ?L2) \<union> (?R1 - ?R2)"
       
   285         using ** by auto
       
   286       also have "\<dots> = set_tree (diff t1 t2)"
       
   287       using IHl IHr sp diff.simps[of t1 t2] False by(simp)
       
   288       finally show ?thesis by simp
       
   289     qed
       
   290   qed
       
   291 qed
       
   292 
       
   293 lemma bst_diff: "\<lbrakk> bst t1; bst t2 \<rbrakk> \<Longrightarrow> bst (diff t1 t2)"
       
   294 proof(induction t1 t2 rule: diff.induct)
       
   295   case (1 t1 t2)
       
   296   thus ?case
       
   297     by(fastforce simp: diff.simps[of t1 t2] set_tree_diff split Let_def
       
   298         intro!: bst_join bst_join2 split: tree.split prod.split)
       
   299 qed
       
   300 
       
   301 lemma inv_diff: "\<lbrakk> inv t1; inv t2 \<rbrakk> \<Longrightarrow> inv (diff t1 t2)"
       
   302 proof(induction t1 t2 rule: diff.induct)
       
   303   case (1 t1 t2)
       
   304   thus ?case
       
   305     by(auto simp: diff.simps[of t1 t2] inv_join inv_join2 split_inv Let_def
       
   306         split!: tree.split prod.split dest: inv_Node)
       
   307 qed
       
   308 
       
   309 text \<open>Locale @{locale Set2_BST2_Join} implements locale @{locale Set2}:\<close>
       
   310 
       
   311 sublocale Set2
       
   312 where empty = Leaf and insert = insert and delete = delete and isin = isin
       
   313 and union = union and inter = inter and diff = diff
       
   314 and set = set_tree and invar = "\<lambda>t. inv t \<and> bst t"
       
   315 proof (standard, goal_cases)
       
   316   case 1 show ?case by (simp)
       
   317 next
       
   318   case 2 thus ?case by(simp add: isin_set_tree)
       
   319 next
       
   320   case 3 thus ?case by (simp add: set_tree_insert)
       
   321 next
       
   322   case 4 thus ?case by (simp add: set_tree_delete)
       
   323 next
       
   324   case 5 thus ?case by (simp add: inv_Leaf)
       
   325 next
       
   326   case 6 thus ?case by (simp add: bst_insert inv_insert)
       
   327 next
       
   328   case 7 thus ?case by (simp add: bst_delete inv_delete)
       
   329 next
       
   330   case 8 thus ?case by(simp add: set_tree_union)
       
   331 next
       
   332   case 9 thus ?case by(simp add: set_tree_inter)
       
   333 next
       
   334   case 10 thus ?case by(simp add: set_tree_diff)
       
   335 next
       
   336   case 11 thus ?case by (simp add: bst_union inv_union)
       
   337 next
       
   338   case 12 thus ?case by (simp add: bst_inter inv_inter)
       
   339 next
       
   340   case 13 thus ?case by (simp add: bst_diff inv_diff)
       
   341 qed
       
   342 
       
   343 end
       
   344 
       
   345 interpretation unbal: Set2_BST2_Join
       
   346 where join = "\<lambda>l x r. Node () l x r" and inv = "\<lambda>t. True"
       
   347 proof (standard, goal_cases)
       
   348   case 1 show ?case by simp
       
   349 next
       
   350   case 2 thus ?case by simp
       
   351 next
       
   352   case 3 thus ?case by simp
       
   353 next
       
   354   case 4 thus ?case by simp
       
   355 next
       
   356   case 5 thus ?case by simp
       
   357 qed
       
   358 
       
   359 end