| author | nipkow | 
| Fri, 28 Feb 2014 18:09:37 +0100 | |
| changeset 55807 | fd31d0e70eb8 | 
| parent 55378 | e61e023c9faf | 
| child 58249 | 180f1b3508ed | 
| permissions | -rw-r--r-- | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
28819diff
changeset | 1 | (* Title: HOL/Statespace/DistinctTreeProver.thy | 
| 25171 | 2 | Author: Norbert Schirmer, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 5 | header {* Distinctness of Names in a Binary Tree \label{sec:DistinctTreeProver}*}
 | |
| 6 | ||
| 7 | theory DistinctTreeProver | |
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 11 | text {* A state space manages a set of (abstract) names and assumes
 | |
| 12 | that the names are distinct. The names are stored as parameters of a | |
| 13 | locale and distinctness as an assumption. The most common request is | |
| 14 | to proof distinctness of two given names. We maintain the names in a | |
| 15 | balanced binary tree and formulate a predicate that all nodes in the | |
| 16 | tree have distinct names. This setup leads to logarithmic certificates. | |
| 17 | *} | |
| 18 | ||
| 19 | subsection {* The Binary Tree *}
 | |
| 20 | ||
| 21 | datatype 'a tree = Node "'a tree" 'a bool "'a tree" | Tip | |
| 22 | ||
| 23 | ||
| 24 | text {* The boolean flag in the node marks the content of the node as
 | |
| 25 | deleted, without having to build a new tree. We prefer the boolean | |
| 26 | flag to an option type, so that the ML-layer can still use the node | |
| 27 | content to facilitate binary search in the tree. The ML code keeps the | |
| 28 | nodes sorted using the term order. We do not have to push ordering to | |
| 29 | the HOL level. *} | |
| 30 | ||
| 31 | subsection {* Distinctness of Nodes *}
 | |
| 32 | ||
| 33 | ||
| 38838 | 34 | primrec set_of :: "'a tree \<Rightarrow> 'a set" | 
| 35 | where | |
| 36 |   "set_of Tip = {}"
 | |
| 37 | | "set_of (Node l x d r) = (if d then {} else {x}) \<union> set_of l \<union> set_of r"
 | |
| 25171 | 38 | |
| 38838 | 39 | primrec all_distinct :: "'a tree \<Rightarrow> bool" | 
| 40 | where | |
| 41 | "all_distinct Tip = True" | |
| 42 | | "all_distinct (Node l x d r) = | |
| 43 | ((d \<or> (x \<notin> set_of l \<and> x \<notin> set_of r)) \<and> | |
| 44 |       set_of l \<inter> set_of r = {} \<and>
 | |
| 45 | all_distinct l \<and> all_distinct r)" | |
| 25171 | 46 | |
| 47 | text {* Given a binary tree @{term "t"} for which 
 | |
| 48 | @{const all_distinct} holds, given two different nodes contained in the tree,
 | |
| 49 | we want to write a ML function that generates a logarithmic | |
| 50 | certificate that the content of the nodes is distinct. We use the | |
| 51 | following lemmas to achieve this. *} | |
| 52 | ||
| 45355 | 53 | lemma all_distinct_left: "all_distinct (Node l x b r) \<Longrightarrow> all_distinct l" | 
| 25171 | 54 | by simp | 
| 55 | ||
| 56 | lemma all_distinct_right: "all_distinct (Node l x b r) \<Longrightarrow> all_distinct r" | |
| 57 | by simp | |
| 58 | ||
| 45358 | 59 | lemma distinct_left: "all_distinct (Node l x False r) \<Longrightarrow> y \<in> set_of l \<Longrightarrow> x \<noteq> y" | 
| 25171 | 60 | by auto | 
| 61 | ||
| 45358 | 62 | lemma distinct_right: "all_distinct (Node l x False r) \<Longrightarrow> y \<in> set_of r \<Longrightarrow> x \<noteq> y" | 
| 25171 | 63 | by auto | 
| 64 | ||
| 45358 | 65 | lemma distinct_left_right: | 
| 66 | "all_distinct (Node l z b r) \<Longrightarrow> x \<in> set_of l \<Longrightarrow> y \<in> set_of r \<Longrightarrow> x \<noteq> y" | |
| 25171 | 67 | by auto | 
| 68 | ||
| 69 | lemma in_set_root: "x \<in> set_of (Node l x False r)" | |
| 70 | by simp | |
| 71 | ||
| 72 | lemma in_set_left: "y \<in> set_of l \<Longrightarrow> y \<in> set_of (Node l x False r)" | |
| 73 | by simp | |
| 74 | ||
| 75 | lemma in_set_right: "y \<in> set_of r \<Longrightarrow> y \<in> set_of (Node l x False r)" | |
| 76 | by simp | |
| 77 | ||
| 78 | lemma swap_neq: "x \<noteq> y \<Longrightarrow> y \<noteq> x" | |
| 79 | by blast | |
| 80 | ||
| 81 | lemma neq_to_eq_False: "x\<noteq>y \<Longrightarrow> (x=y)\<equiv>False" | |
| 82 | by simp | |
| 83 | ||
| 84 | subsection {* Containment of Trees *}
 | |
| 85 | ||
| 86 | text {* When deriving a state space from other ones, we create a new
 | |
| 87 | name tree which contains all the names of the parent state spaces and | |
| 45358 | 88 | assume the predicate @{const all_distinct}. We then prove that the new
 | 
| 89 | locale interprets all parent locales. Hence we have to show that the | |
| 90 | new distinctness assumption on all names implies the distinctness | |
| 25171 | 91 | assumptions of the parent locales. This proof is implemented in ML. We | 
| 92 | do this efficiently by defining a kind of containment check of trees | |
| 45358 | 93 | by ``subtraction''. We subtract the parent tree from the new tree. If | 
| 94 | this succeeds we know that @{const all_distinct} of the new tree
 | |
| 95 | implies @{const all_distinct} of the parent tree.  The resulting
 | |
| 96 | certificate is of the order @{term "n * log(m)"} where @{term "n"} is
 | |
| 97 | the size of the (smaller) parent tree and @{term "m"} the size of the
 | |
| 98 | (bigger) new tree. *} | |
| 25171 | 99 | |
| 100 | ||
| 38838 | 101 | primrec delete :: "'a \<Rightarrow> 'a tree \<Rightarrow> 'a tree option" | 
| 102 | where | |
| 103 | "delete x Tip = None" | |
| 104 | | "delete x (Node l y d r) = (case delete x l of | |
| 105 | Some l' \<Rightarrow> | |
| 106 | (case delete x r of | |
| 107 | Some r' \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r') | |
| 108 | | None \<Rightarrow> Some (Node l' y (d \<or> (x=y)) r)) | |
| 109 | | None \<Rightarrow> | |
| 45358 | 110 | (case delete x r of | 
| 38838 | 111 | Some r' \<Rightarrow> Some (Node l y (d \<or> (x=y)) r') | 
| 112 | | None \<Rightarrow> if x=y \<and> \<not>d then Some (Node l y True r) | |
| 113 | else None))" | |
| 25171 | 114 | |
| 115 | ||
| 45358 | 116 | lemma delete_Some_set_of: "delete x t = Some t' \<Longrightarrow> set_of t' \<subseteq> set_of t" | 
| 117 | proof (induct t arbitrary: t') | |
| 25171 | 118 | case Tip thus ?case by simp | 
| 119 | next | |
| 120 | case (Node l y d r) | |
| 25364 | 121 | have del: "delete x (Node l y d r) = Some t'" by fact | 
| 25171 | 122 | show ?case | 
| 123 | proof (cases "delete x l") | |
| 124 | case (Some l') | |
| 125 | note x_l_Some = this | |
| 126 | with Node.hyps | |
| 127 | have l'_l: "set_of l' \<subseteq> set_of l" | |
| 128 | by simp | |
| 129 | show ?thesis | |
| 130 | proof (cases "delete x r") | |
| 131 | case (Some r') | |
| 132 | with Node.hyps | |
| 133 | have "set_of r' \<subseteq> set_of r" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 134 | by simp | 
| 25171 | 135 | with l'_l Some x_l_Some del | 
| 136 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 137 | by (auto split: split_if_asm) | 
| 25171 | 138 | next | 
| 139 | case None | |
| 140 | with l'_l Some x_l_Some del | |
| 141 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 142 | by (fastforce split: split_if_asm) | 
| 25171 | 143 | qed | 
| 144 | next | |
| 145 | case None | |
| 146 | note x_l_None = this | |
| 147 | show ?thesis | |
| 148 | proof (cases "delete x r") | |
| 149 | case (Some r') | |
| 150 | with Node.hyps | |
| 151 | have "set_of r' \<subseteq> set_of r" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 152 | by simp | 
| 25171 | 153 | with Some x_l_None del | 
| 154 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 155 | by (fastforce split: split_if_asm) | 
| 25171 | 156 | next | 
| 157 | case None | |
| 158 | with x_l_None del | |
| 159 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 160 | by (fastforce split: split_if_asm) | 
| 25171 | 161 | qed | 
| 162 | qed | |
| 163 | qed | |
| 164 | ||
| 45355 | 165 | lemma delete_Some_all_distinct: | 
| 45358 | 166 | "delete x t = Some t' \<Longrightarrow> all_distinct t \<Longrightarrow> all_distinct t'" | 
| 167 | proof (induct t arbitrary: t') | |
| 25171 | 168 | case Tip thus ?case by simp | 
| 169 | next | |
| 170 | case (Node l y d r) | |
| 25364 | 171 | have del: "delete x (Node l y d r) = Some t'" by fact | 
| 172 | have "all_distinct (Node l y d r)" by fact | |
| 25171 | 173 | then obtain | 
| 174 | dist_l: "all_distinct l" and | |
| 175 | dist_r: "all_distinct r" and | |
| 176 | d: "d \<or> (y \<notin> set_of l \<and> y \<notin> set_of r)" and | |
| 177 |     dist_l_r: "set_of l \<inter> set_of r = {}"
 | |
| 178 | by auto | |
| 179 | show ?case | |
| 180 | proof (cases "delete x l") | |
| 181 | case (Some l') | |
| 182 | note x_l_Some = this | |
| 183 | from Node.hyps (1) [OF Some dist_l] | |
| 184 | have dist_l': "all_distinct l'" | |
| 185 | by simp | |
| 186 | from delete_Some_set_of [OF x_l_Some] | |
| 187 | have l'_l: "set_of l' \<subseteq> set_of l". | |
| 188 | show ?thesis | |
| 189 | proof (cases "delete x r") | |
| 190 | case (Some r') | |
| 191 | from Node.hyps (2) [OF Some dist_r] | |
| 192 | have dist_r': "all_distinct r'" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 193 | by simp | 
| 25171 | 194 | from delete_Some_set_of [OF Some] | 
| 195 | have "set_of r' \<subseteq> set_of r". | |
| 196 | ||
| 197 | with dist_l' dist_r' l'_l Some x_l_Some del d dist_l_r | |
| 198 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 199 | by fastforce | 
| 25171 | 200 | next | 
| 201 | case None | |
| 202 | with l'_l dist_l' x_l_Some del d dist_l_r dist_r | |
| 203 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 204 | by fastforce | 
| 25171 | 205 | qed | 
| 206 | next | |
| 207 | case None | |
| 208 | note x_l_None = this | |
| 209 | show ?thesis | |
| 210 | proof (cases "delete x r") | |
| 211 | case (Some r') | |
| 212 | with Node.hyps (2) [OF Some dist_r] | |
| 213 | have dist_r': "all_distinct r'" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 214 | by simp | 
| 25171 | 215 | from delete_Some_set_of [OF Some] | 
| 216 | have "set_of r' \<subseteq> set_of r". | |
| 217 | with Some dist_r' x_l_None del dist_l d dist_l_r | |
| 218 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 219 | by fastforce | 
| 25171 | 220 | next | 
| 221 | case None | |
| 222 | with x_l_None del dist_l dist_r d dist_l_r | |
| 223 | show ?thesis | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42287diff
changeset | 224 | by (fastforce split: split_if_asm) | 
| 25171 | 225 | qed | 
| 226 | qed | |
| 227 | qed | |
| 228 | ||
| 229 | lemma delete_None_set_of_conv: "delete x t = None = (x \<notin> set_of t)" | |
| 230 | proof (induct t) | |
| 231 | case Tip thus ?case by simp | |
| 232 | next | |
| 233 | case (Node l y d r) | |
| 234 | thus ?case | |
| 235 | by (auto split: option.splits) | |
| 236 | qed | |
| 237 | ||
| 238 | lemma delete_Some_x_set_of: | |
| 45358 | 239 | "delete x t = Some t' \<Longrightarrow> x \<in> set_of t \<and> x \<notin> set_of t'" | 
| 240 | proof (induct t arbitrary: t') | |
| 25171 | 241 | case Tip thus ?case by simp | 
| 242 | next | |
| 243 | case (Node l y d r) | |
| 25364 | 244 | have del: "delete x (Node l y d r) = Some t'" by fact | 
| 25171 | 245 | show ?case | 
| 246 | proof (cases "delete x l") | |
| 247 | case (Some l') | |
| 248 | note x_l_Some = this | |
| 249 | from Node.hyps (1) [OF Some] | |
| 250 | obtain x_l: "x \<in> set_of l" "x \<notin> set_of l'" | |
| 251 | by simp | |
| 252 | show ?thesis | |
| 253 | proof (cases "delete x r") | |
| 254 | case (Some r') | |
| 255 | from Node.hyps (2) [OF Some] | |
| 256 | obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 257 | by simp | 
| 25171 | 258 | from x_r x_l Some x_l_Some del | 
| 259 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 260 | by (clarsimp split: split_if_asm) | 
| 25171 | 261 | next | 
| 262 | case None | |
| 263 | then have "x \<notin> set_of r" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 264 | by (simp add: delete_None_set_of_conv) | 
| 25171 | 265 | with x_l None x_l_Some del | 
| 266 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 267 | by (clarsimp split: split_if_asm) | 
| 25171 | 268 | qed | 
| 269 | next | |
| 270 | case None | |
| 271 | note x_l_None = this | |
| 272 | then have x_notin_l: "x \<notin> set_of l" | |
| 273 | by (simp add: delete_None_set_of_conv) | |
| 274 | show ?thesis | |
| 275 | proof (cases "delete x r") | |
| 276 | case (Some r') | |
| 277 | from Node.hyps (2) [OF Some] | |
| 278 | obtain x_r: "x \<in> set_of r" "x \<notin> set_of r'" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 279 | by simp | 
| 25171 | 280 | from x_r x_notin_l Some x_l_None del | 
| 281 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 282 | by (clarsimp split: split_if_asm) | 
| 25171 | 283 | next | 
| 284 | case None | |
| 285 | then have "x \<notin> set_of r" | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 286 | by (simp add: delete_None_set_of_conv) | 
| 25171 | 287 | with None x_l_None x_notin_l del | 
| 288 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 289 | by (clarsimp split: split_if_asm) | 
| 25171 | 290 | qed | 
| 291 | qed | |
| 292 | qed | |
| 293 | ||
| 294 | ||
| 38838 | 295 | primrec subtract :: "'a tree \<Rightarrow> 'a tree \<Rightarrow> 'a tree option" | 
| 296 | where | |
| 297 | "subtract Tip t = Some t" | |
| 298 | | "subtract (Node l x b r) t = | |
| 299 | (case delete x t of | |
| 300 | Some t' \<Rightarrow> (case subtract l t' of | |
| 301 | Some t'' \<Rightarrow> subtract r t'' | |
| 302 | | None \<Rightarrow> None) | |
| 303 | | None \<Rightarrow> None)" | |
| 25171 | 304 | |
| 305 | lemma subtract_Some_set_of_res: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 306 | "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t \<subseteq> set_of t\<^sub>2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 307 | proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t) | 
| 25171 | 308 | case Tip thus ?case by simp | 
| 309 | next | |
| 310 | case (Node l x b r) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 311 | have sub: "subtract (Node l x b r) t\<^sub>2 = Some t" by fact | 
| 25171 | 312 | show ?case | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 313 | proof (cases "delete x t\<^sub>2") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 314 | case (Some t\<^sub>2') | 
| 25171 | 315 | note del_x_Some = this | 
| 316 | from delete_Some_set_of [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 317 | have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" . | 
| 25171 | 318 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 319 | proof (cases "subtract l t\<^sub>2'") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 320 | case (Some t\<^sub>2'') | 
| 25171 | 321 | note sub_l_Some = this | 
| 322 | from Node.hyps (1) [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 323 | have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 324 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 325 | proof (cases "subtract r t\<^sub>2''") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 326 | case (Some t\<^sub>2''') | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 327 | from Node.hyps (2) [OF Some ] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 328 | have "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''" . | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 329 | with Some sub_l_Some del_x_Some sub t2''_t2' t2'_t2 | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 330 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 331 | by simp | 
| 25171 | 332 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 333 | case None | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 334 | with del_x_Some sub_l_Some sub | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 335 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 336 | by simp | 
| 25171 | 337 | qed | 
| 338 | next | |
| 339 | case None | |
| 340 | with del_x_Some sub | |
| 341 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 342 | by simp | 
| 25171 | 343 | qed | 
| 344 | next | |
| 345 | case None | |
| 346 | with sub show ?thesis by simp | |
| 347 | qed | |
| 348 | qed | |
| 349 | ||
| 350 | lemma subtract_Some_set_of: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 351 | "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<subseteq> set_of t\<^sub>2" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 352 | proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t) | 
| 25171 | 353 | case Tip thus ?case by simp | 
| 354 | next | |
| 355 | case (Node l x d r) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 356 | have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact | 
| 25171 | 357 | show ?case | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 358 | proof (cases "delete x t\<^sub>2") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 359 | case (Some t\<^sub>2') | 
| 25171 | 360 | note del_x_Some = this | 
| 361 | from delete_Some_set_of [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 362 | have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" . | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 363 | from delete_None_set_of_conv [of x t\<^sub>2] Some | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 364 | have x_t2: "x \<in> set_of t\<^sub>2" | 
| 25171 | 365 | by simp | 
| 366 | show ?thesis | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 367 | proof (cases "subtract l t\<^sub>2'") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 368 | case (Some t\<^sub>2'') | 
| 25171 | 369 | note sub_l_Some = this | 
| 370 | from Node.hyps (1) [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 371 | have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 372 | from subtract_Some_set_of_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 373 | have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 374 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 375 | proof (cases "subtract r t\<^sub>2''") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 376 | case (Some t\<^sub>2''') | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 377 | from Node.hyps (2) [OF Some ] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 378 | have r_t\<^sub>2'': "set_of r \<subseteq> set_of t\<^sub>2''" . | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 379 | from Some sub_l_Some del_x_Some sub r_t\<^sub>2'' l_t2' t2'_t2 t2''_t2' x_t2 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 380 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 381 | by auto | 
| 25171 | 382 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 383 | case None | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 384 | with del_x_Some sub_l_Some sub | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 385 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 386 | by simp | 
| 25171 | 387 | qed | 
| 388 | next | |
| 389 | case None | |
| 390 | with del_x_Some sub | |
| 391 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 392 | by simp | 
| 25171 | 393 | qed | 
| 394 | next | |
| 395 | case None | |
| 396 | with sub show ?thesis by simp | |
| 397 | qed | |
| 398 | qed | |
| 399 | ||
| 400 | lemma subtract_Some_all_distinct_res: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 401 | "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 402 | proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t) | 
| 25171 | 403 | case Tip thus ?case by simp | 
| 404 | next | |
| 405 | case (Node l x d r) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 406 | have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 407 | have dist_t2: "all_distinct t\<^sub>2" by fact | 
| 25171 | 408 | show ?case | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 409 | proof (cases "delete x t\<^sub>2") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 410 | case (Some t\<^sub>2') | 
| 25171 | 411 | note del_x_Some = this | 
| 412 | from delete_Some_all_distinct [OF Some dist_t2] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 413 | have dist_t2': "all_distinct t\<^sub>2'" . | 
| 25171 | 414 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 415 | proof (cases "subtract l t\<^sub>2'") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 416 | case (Some t\<^sub>2'') | 
| 25171 | 417 | note sub_l_Some = this | 
| 418 | from Node.hyps (1) [OF Some dist_t2'] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 419 | have dist_t2'': "all_distinct t\<^sub>2''" . | 
| 25171 | 420 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 421 | proof (cases "subtract r t\<^sub>2''") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 422 | case (Some t\<^sub>2''') | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 423 | from Node.hyps (2) [OF Some dist_t2''] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 424 | have dist_t2''': "all_distinct t\<^sub>2'''" . | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 425 | from Some sub_l_Some del_x_Some sub | 
| 25171 | 426 | dist_t2''' | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 427 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 428 | by simp | 
| 25171 | 429 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 430 | case None | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 431 | with del_x_Some sub_l_Some sub | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 432 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 433 | by simp | 
| 25171 | 434 | qed | 
| 435 | next | |
| 436 | case None | |
| 437 | with del_x_Some sub | |
| 438 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 439 | by simp | 
| 25171 | 440 | qed | 
| 441 | next | |
| 442 | case None | |
| 443 | with sub show ?thesis by simp | |
| 444 | qed | |
| 445 | qed | |
| 446 | ||
| 447 | ||
| 448 | lemma subtract_Some_dist_res: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 449 |   "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> set_of t\<^sub>1 \<inter> set_of t = {}"
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 450 | proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t) | 
| 25171 | 451 | case Tip thus ?case by simp | 
| 452 | next | |
| 453 | case (Node l x d r) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 454 | have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact | 
| 25171 | 455 | show ?case | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 456 | proof (cases "delete x t\<^sub>2") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 457 | case (Some t\<^sub>2') | 
| 25171 | 458 | note del_x_Some = this | 
| 459 | from delete_Some_x_set_of [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 460 | obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'" | 
| 25171 | 461 | by simp | 
| 462 | from delete_Some_set_of [OF Some] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 463 | have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" . | 
| 25171 | 464 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 465 | proof (cases "subtract l t\<^sub>2'") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 466 | case (Some t\<^sub>2'') | 
| 25171 | 467 | note sub_l_Some = this | 
| 468 | from Node.hyps (1) [OF Some ] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 469 |       have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
 | 
| 25171 | 470 | from subtract_Some_set_of_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 471 | have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 472 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 473 | proof (cases "subtract r t\<^sub>2''") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 474 | case (Some t\<^sub>2''') | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 475 | from Node.hyps (2) [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 476 |         have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}" .
 | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 477 | from subtract_Some_set_of_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 478 | have t2'''_t2'': "set_of t\<^sub>2''' \<subseteq> set_of t\<^sub>2''". | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 479 | |
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 480 | from Some sub_l_Some del_x_Some sub t2'''_t2'' dist_l_t2'' dist_r_t2''' | 
| 25171 | 481 | t2''_t2' t2'_t2 x_not_t2' | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 482 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 483 | by auto | 
| 25171 | 484 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 485 | case None | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 486 | with del_x_Some sub_l_Some sub | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 487 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 488 | by simp | 
| 25171 | 489 | qed | 
| 490 | next | |
| 491 | case None | |
| 492 | with del_x_Some sub | |
| 493 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 494 | by simp | 
| 25171 | 495 | qed | 
| 496 | next | |
| 497 | case None | |
| 498 | with sub show ?thesis by simp | |
| 499 | qed | |
| 500 | qed | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 501 | |
| 25171 | 502 | lemma subtract_Some_all_distinct: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 503 | "subtract t\<^sub>1 t\<^sub>2 = Some t \<Longrightarrow> all_distinct t\<^sub>2 \<Longrightarrow> all_distinct t\<^sub>1" | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 504 | proof (induct t\<^sub>1 arbitrary: t\<^sub>2 t) | 
| 25171 | 505 | case Tip thus ?case by simp | 
| 506 | next | |
| 507 | case (Node l x d r) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 508 | have sub: "subtract (Node l x d r) t\<^sub>2 = Some t" by fact | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 509 | have dist_t2: "all_distinct t\<^sub>2" by fact | 
| 25171 | 510 | show ?case | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 511 | proof (cases "delete x t\<^sub>2") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 512 | case (Some t\<^sub>2') | 
| 25171 | 513 | note del_x_Some = this | 
| 514 | from delete_Some_all_distinct [OF Some dist_t2 ] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 515 | have dist_t2': "all_distinct t\<^sub>2'" . | 
| 25171 | 516 | from delete_Some_set_of [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 517 | have t2'_t2: "set_of t\<^sub>2' \<subseteq> set_of t\<^sub>2" . | 
| 25171 | 518 | from delete_Some_x_set_of [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 519 | obtain x_t2: "x \<in> set_of t\<^sub>2" and x_not_t2': "x \<notin> set_of t\<^sub>2'" | 
| 25171 | 520 | by simp | 
| 521 | ||
| 522 | show ?thesis | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 523 | proof (cases "subtract l t\<^sub>2'") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 524 | case (Some t\<^sub>2'') | 
| 25171 | 525 | note sub_l_Some = this | 
| 526 | from Node.hyps (1) [OF Some dist_t2' ] | |
| 527 | have dist_l: "all_distinct l" . | |
| 528 | from subtract_Some_all_distinct_res [OF Some dist_t2'] | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 529 | have dist_t2'': "all_distinct t\<^sub>2''" . | 
| 25171 | 530 | from subtract_Some_set_of [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 531 | have l_t2': "set_of l \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 532 | from subtract_Some_set_of_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 533 | have t2''_t2': "set_of t\<^sub>2'' \<subseteq> set_of t\<^sub>2'" . | 
| 25171 | 534 | from subtract_Some_dist_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 535 |       have dist_l_t2'': "set_of l \<inter> set_of t\<^sub>2'' = {}".
 | 
| 25171 | 536 | show ?thesis | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 537 | proof (cases "subtract r t\<^sub>2''") | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 538 | case (Some t\<^sub>2''') | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 539 | from Node.hyps (2) [OF Some dist_t2''] | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 540 | have dist_r: "all_distinct r" . | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 541 | from subtract_Some_set_of [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 542 | have r_t2'': "set_of r \<subseteq> set_of t\<^sub>2''" . | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 543 | from subtract_Some_dist_res [OF Some] | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
48891diff
changeset | 544 |         have dist_r_t2''': "set_of r \<inter> set_of t\<^sub>2''' = {}".
 | 
| 25171 | 545 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 546 | from dist_l dist_r Some sub_l_Some del_x_Some r_t2'' l_t2' x_t2 x_not_t2' | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 547 | t2''_t2' dist_l_t2'' dist_r_t2''' | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 548 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 549 | by auto | 
| 25171 | 550 | next | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 551 | case None | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 552 | with del_x_Some sub_l_Some sub | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 553 | show ?thesis | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 554 | by simp | 
| 25171 | 555 | qed | 
| 556 | next | |
| 557 | case None | |
| 558 | with del_x_Some sub | |
| 559 | show ?thesis | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
32740diff
changeset | 560 | by simp | 
| 25171 | 561 | qed | 
| 562 | next | |
| 563 | case None | |
| 564 | with sub show ?thesis by simp | |
| 565 | qed | |
| 566 | qed | |
| 567 | ||
| 568 | ||
| 569 | lemma delete_left: | |
| 570 | assumes dist: "all_distinct (Node l y d r)" | |
| 571 | assumes del_l: "delete x l = Some l'" | |
| 572 | shows "delete x (Node l y d r) = Some (Node l' y d r)" | |
| 573 | proof - | |
| 574 | from delete_Some_x_set_of [OF del_l] | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 575 | obtain x: "x \<in> set_of l" | 
| 25171 | 576 | by simp | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 577 | with dist | 
| 25171 | 578 | have "delete x r = None" | 
| 579 | by (cases "delete x r") (auto dest:delete_Some_x_set_of) | |
| 580 | ||
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 581 | with x | 
| 25171 | 582 | show ?thesis | 
| 583 | using del_l dist | |
| 584 | by (auto split: option.splits) | |
| 585 | qed | |
| 586 | ||
| 587 | lemma delete_right: | |
| 588 | assumes dist: "all_distinct (Node l y d r)" | |
| 589 | assumes del_r: "delete x r = Some r'" | |
| 590 | shows "delete x (Node l y d r) = Some (Node l y d r')" | |
| 591 | proof - | |
| 592 | from delete_Some_x_set_of [OF del_r] | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 593 | obtain x: "x \<in> set_of r" | 
| 25171 | 594 | by simp | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 595 | with dist | 
| 25171 | 596 | have "delete x l = None" | 
| 597 | by (cases "delete x l") (auto dest:delete_Some_x_set_of) | |
| 598 | ||
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 599 | with x | 
| 25171 | 600 | show ?thesis | 
| 601 | using del_r dist | |
| 602 | by (auto split: option.splits) | |
| 603 | qed | |
| 604 | ||
| 605 | lemma delete_root: | |
| 606 | assumes dist: "all_distinct (Node l x False r)" | |
| 607 | shows "delete x (Node l x False r) = Some (Node l x True r)" | |
| 608 | proof - | |
| 609 | from dist have "delete x r = None" | |
| 610 | by (cases "delete x r") (auto dest:delete_Some_x_set_of) | |
| 611 | moreover | |
| 612 | from dist have "delete x l = None" | |
| 613 | by (cases "delete x l") (auto dest:delete_Some_x_set_of) | |
| 614 | ultimately show ?thesis | |
| 615 | using dist | |
| 616 | by (auto split: option.splits) | |
| 617 | qed | |
| 618 | ||
| 619 | lemma subtract_Node: | |
| 620 | assumes del: "delete x t = Some t'" | |
| 621 | assumes sub_l: "subtract l t' = Some t''" | |
| 622 | assumes sub_r: "subtract r t'' = Some t'''" | |
| 623 | shows "subtract (Node l x False r) t = Some t'''" | |
| 624 | using del sub_l sub_r | |
| 625 | by simp | |
| 626 | ||
| 627 | lemma subtract_Tip: "subtract Tip t = Some t" | |
| 628 | by simp | |
| 629 | ||
| 630 | text {* Now we have all the theorems in place that are needed for the
 | |
| 631 | certificate generating ML functions. *} | |
| 632 | ||
| 48891 | 633 | ML_file "distinct_tree_prover.ML" | 
| 25171 | 634 | |
| 635 | end |