| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 19 Jan 2011 18:56:28 +0100 | |
| changeset 2683 | 42c0d011a177 | 
| parent 2679 | e003e5e36bae | 
| child 2685 | 1df873b63cb2 | 
| permissions | -rw-r--r-- | 
| 1062 | 1 | (* Title: Nominal2_Base | 
| 2 | Authors: Brian Huffman, Christian Urban | |
| 3 | ||
| 4 | Basic definitions and lemma infrastructure for | |
| 5 | Nominal Isabelle. | |
| 6 | *) | |
| 7 | theory Nominal2_Base | |
| 2635 
64b4cb2c2bf8
simple cases for string rule inductions
 Christian Urban <urbanc@in.tum.de> parents: 
2632diff
changeset | 8 | imports Main | 
| 
64b4cb2c2bf8
simple cases for string rule inductions
 Christian Urban <urbanc@in.tum.de> parents: 
2632diff
changeset | 9 | "~~/src/HOL/Library/Infinite_Set" | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 10 | "~~/src/HOL/Quotient_Examples/FSet" | 
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1774diff
changeset | 11 | uses ("nominal_library.ML")
 | 
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 12 |      ("nominal_atoms.ML")
 | 
| 1062 | 13 | begin | 
| 14 | ||
| 15 | section {* Atoms and Sorts *}
 | |
| 16 | ||
| 17 | text {* A simple implementation for atom_sorts is strings. *}
 | |
| 18 | (* types atom_sort = string *) | |
| 19 | ||
| 20 | text {* To deal with Church-like binding we use trees of  
 | |
| 21 | strings as sorts. *} | |
| 22 | ||
| 23 | datatype atom_sort = Sort "string" "atom_sort list" | |
| 24 | ||
| 25 | datatype atom = Atom atom_sort nat | |
| 26 | ||
| 27 | ||
| 28 | text {* Basic projection function. *}
 | |
| 29 | ||
| 30 | primrec | |
| 31 | sort_of :: "atom \<Rightarrow> atom_sort" | |
| 32 | where | |
| 33 | "sort_of (Atom s i) = s" | |
| 34 | ||
| 1930 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 35 | primrec | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 36 | nat_of :: "atom \<Rightarrow> nat" | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 37 | where | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 38 | "nat_of (Atom s n) = n" | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 39 | |
| 1062 | 40 | |
| 41 | text {* There are infinitely many atoms of each sort. *}
 | |
| 42 | lemma INFM_sort_of_eq: | |
| 43 | shows "INFM a. sort_of a = s" | |
| 44 | proof - | |
| 45 | have "INFM i. sort_of (Atom s i) = s" by simp | |
| 46 | moreover have "inj (Atom s)" by (simp add: inj_on_def) | |
| 47 | ultimately show "INFM a. sort_of a = s" by (rule INFM_inj) | |
| 48 | qed | |
| 49 | ||
| 50 | lemma infinite_sort_of_eq: | |
| 51 |   shows "infinite {a. sort_of a = s}"
 | |
| 52 | using INFM_sort_of_eq unfolding INFM_iff_infinite . | |
| 53 | ||
| 54 | lemma atom_infinite [simp]: | |
| 55 | shows "infinite (UNIV :: atom set)" | |
| 56 | using subset_UNIV infinite_sort_of_eq | |
| 57 | by (rule infinite_super) | |
| 58 | ||
| 59 | lemma obtain_atom: | |
| 60 | fixes X :: "atom set" | |
| 61 | assumes X: "finite X" | |
| 62 | obtains a where "a \<notin> X" "sort_of a = s" | |
| 63 | proof - | |
| 64 | from X have "MOST a. a \<notin> X" | |
| 65 | unfolding MOST_iff_cofinite by simp | |
| 66 | with INFM_sort_of_eq | |
| 67 | have "INFM a. sort_of a = s \<and> a \<notin> X" | |
| 68 | by (rule INFM_conjI) | |
| 69 | then obtain a where "a \<notin> X" "sort_of a = s" | |
| 70 | by (auto elim: INFM_E) | |
| 71 | then show ?thesis .. | |
| 72 | qed | |
| 73 | ||
| 1930 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 74 | lemma atom_components_eq_iff: | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 75 | fixes a b :: atom | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 76 | shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b" | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 77 | by (induct a, induct b, simp) | 
| 
f189cf2c0987
moved some lemmas into the right places
 Christian Urban <urbanc@in.tum.de> parents: 
1879diff
changeset | 78 | |
| 1062 | 79 | section {* Sort-Respecting Permutations *}
 | 
| 80 | ||
| 81 | typedef perm = | |
| 82 |   "{f. bij f \<and> finite {a. f a \<noteq> a} \<and> (\<forall>a. sort_of (f a) = sort_of a)}"
 | |
| 83 | proof | |
| 84 | show "id \<in> ?perm" by simp | |
| 85 | qed | |
| 86 | ||
| 87 | lemma permI: | |
| 88 | assumes "bij f" and "MOST x. f x = x" and "\<And>a. sort_of (f a) = sort_of a" | |
| 89 | shows "f \<in> perm" | |
| 90 | using assms unfolding perm_def MOST_iff_cofinite by simp | |
| 91 | ||
| 92 | lemma perm_is_bij: "f \<in> perm \<Longrightarrow> bij f" | |
| 93 | unfolding perm_def by simp | |
| 94 | ||
| 95 | lemma perm_is_finite: "f \<in> perm \<Longrightarrow> finite {a. f a \<noteq> a}"
 | |
| 96 | unfolding perm_def by simp | |
| 97 | ||
| 98 | lemma perm_is_sort_respecting: "f \<in> perm \<Longrightarrow> sort_of (f a) = sort_of a" | |
| 99 | unfolding perm_def by simp | |
| 100 | ||
| 101 | lemma perm_MOST: "f \<in> perm \<Longrightarrow> MOST x. f x = x" | |
| 102 | unfolding perm_def MOST_iff_cofinite by simp | |
| 103 | ||
| 104 | lemma perm_id: "id \<in> perm" | |
| 105 | unfolding perm_def by simp | |
| 106 | ||
| 107 | lemma perm_comp: | |
| 108 | assumes f: "f \<in> perm" and g: "g \<in> perm" | |
| 109 | shows "(f \<circ> g) \<in> perm" | |
| 110 | apply (rule permI) | |
| 111 | apply (rule bij_comp) | |
| 112 | apply (rule perm_is_bij [OF g]) | |
| 113 | apply (rule perm_is_bij [OF f]) | |
| 114 | apply (rule MOST_rev_mp [OF perm_MOST [OF g]]) | |
| 115 | apply (rule MOST_rev_mp [OF perm_MOST [OF f]]) | |
| 116 | apply (simp) | |
| 117 | apply (simp add: perm_is_sort_respecting [OF f]) | |
| 118 | apply (simp add: perm_is_sort_respecting [OF g]) | |
| 119 | done | |
| 120 | ||
| 121 | lemma perm_inv: | |
| 122 | assumes f: "f \<in> perm" | |
| 123 | shows "(inv f) \<in> perm" | |
| 124 | apply (rule permI) | |
| 125 | apply (rule bij_imp_bij_inv) | |
| 126 | apply (rule perm_is_bij [OF f]) | |
| 127 | apply (rule MOST_mono [OF perm_MOST [OF f]]) | |
| 128 | apply (erule subst, rule inv_f_f) | |
| 129 | apply (rule bij_is_inj [OF perm_is_bij [OF f]]) | |
| 130 | apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans]) | |
| 131 | apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]]) | |
| 132 | done | |
| 133 | ||
| 134 | lemma bij_Rep_perm: "bij (Rep_perm p)" | |
| 135 | using Rep_perm [of p] unfolding perm_def by simp | |
| 136 | ||
| 137 | lemma finite_Rep_perm: "finite {a. Rep_perm p a \<noteq> a}"
 | |
| 138 | using Rep_perm [of p] unfolding perm_def by simp | |
| 139 | ||
| 140 | lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a" | |
| 141 | using Rep_perm [of p] unfolding perm_def by simp | |
| 142 | ||
| 143 | lemma Rep_perm_ext: | |
| 144 | "Rep_perm p1 = Rep_perm p2 \<Longrightarrow> p1 = p2" | |
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 145 | by (simp add: fun_eq_iff Rep_perm_inject [symmetric]) | 
| 1062 | 146 | |
| 2560 
82e37a4595c7
automated permute_bn functions (raw ones first)
 Christian Urban <urbanc@in.tum.de> parents: 
2507diff
changeset | 147 | instance perm :: size .. | 
| 1062 | 148 | |
| 149 | subsection {* Permutations form a group *}
 | |
| 150 | ||
| 151 | instantiation perm :: group_add | |
| 152 | begin | |
| 153 | ||
| 154 | definition | |
| 155 | "0 = Abs_perm id" | |
| 156 | ||
| 157 | definition | |
| 158 | "- p = Abs_perm (inv (Rep_perm p))" | |
| 159 | ||
| 160 | definition | |
| 161 | "p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)" | |
| 162 | ||
| 163 | definition | |
| 164 | "(p1::perm) - p2 = p1 + - p2" | |
| 165 | ||
| 166 | lemma Rep_perm_0: "Rep_perm 0 = id" | |
| 167 | unfolding zero_perm_def | |
| 168 | by (simp add: Abs_perm_inverse perm_id) | |
| 169 | ||
| 170 | lemma Rep_perm_add: | |
| 171 | "Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2" | |
| 172 | unfolding plus_perm_def | |
| 173 | by (simp add: Abs_perm_inverse perm_comp Rep_perm) | |
| 174 | ||
| 175 | lemma Rep_perm_uminus: | |
| 176 | "Rep_perm (- p) = inv (Rep_perm p)" | |
| 177 | unfolding uminus_perm_def | |
| 178 | by (simp add: Abs_perm_inverse perm_inv Rep_perm) | |
| 179 | ||
| 180 | instance | |
| 181 | apply default | |
| 182 | unfolding Rep_perm_inject [symmetric] | |
| 183 | unfolding minus_perm_def | |
| 184 | unfolding Rep_perm_add | |
| 185 | unfolding Rep_perm_uminus | |
| 186 | unfolding Rep_perm_0 | |
| 187 | by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]]) | |
| 188 | ||
| 189 | end | |
| 190 | ||
| 191 | ||
| 192 | section {* Implementation of swappings *}
 | |
| 193 | ||
| 194 | definition | |
| 195 |   swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')")
 | |
| 196 | where | |
| 197 | "(a \<rightleftharpoons> b) = | |
| 198 | Abs_perm (if sort_of a = sort_of b | |
| 199 | then (\<lambda>c. if a = c then b else if b = c then a else c) | |
| 200 | else id)" | |
| 201 | ||
| 202 | lemma Rep_perm_swap: | |
| 203 | "Rep_perm (a \<rightleftharpoons> b) = | |
| 204 | (if sort_of a = sort_of b | |
| 205 | then (\<lambda>c. if a = c then b else if b = c then a else c) | |
| 206 | else id)" | |
| 207 | unfolding swap_def | |
| 208 | apply (rule Abs_perm_inverse) | |
| 209 | apply (rule permI) | |
| 210 | apply (auto simp add: bij_def inj_on_def surj_def)[1] | |
| 211 | apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]]) | |
| 212 | apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]]) | |
| 213 | apply (simp) | |
| 214 | apply (simp) | |
| 215 | done | |
| 216 | ||
| 217 | lemmas Rep_perm_simps = | |
| 218 | Rep_perm_0 | |
| 219 | Rep_perm_add | |
| 220 | Rep_perm_uminus | |
| 221 | Rep_perm_swap | |
| 222 | ||
| 223 | lemma swap_different_sorts [simp]: | |
| 224 | "sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0" | |
| 225 | by (rule Rep_perm_ext) (simp add: Rep_perm_simps) | |
| 226 | ||
| 227 | lemma swap_cancel: | |
| 2679 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 228 | shows "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 229 | and "(a \<rightleftharpoons> b) + (b \<rightleftharpoons> a) = 0" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 230 | by (rule_tac [!] Rep_perm_ext) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 231 | (simp_all add: Rep_perm_simps fun_eq_iff) | 
| 1062 | 232 | |
| 233 | lemma swap_self [simp]: | |
| 234 | "(a \<rightleftharpoons> a) = 0" | |
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 235 | by (rule Rep_perm_ext, simp add: Rep_perm_simps fun_eq_iff) | 
| 1062 | 236 | |
| 237 | lemma minus_swap [simp]: | |
| 238 | "- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)" | |
| 2679 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 239 | by (rule minus_unique [OF swap_cancel(1)]) | 
| 1062 | 240 | |
| 241 | lemma swap_commute: | |
| 242 | "(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)" | |
| 243 | by (rule Rep_perm_ext) | |
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 244 | (simp add: Rep_perm_swap fun_eq_iff) | 
| 1062 | 245 | |
| 246 | lemma swap_triple: | |
| 247 | assumes "a \<noteq> b" and "c \<noteq> b" | |
| 248 | assumes "sort_of a = sort_of b" "sort_of b = sort_of c" | |
| 249 | shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)" | |
| 250 | using assms | |
| 251 | by (rule_tac Rep_perm_ext) | |
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 252 | (auto simp add: Rep_perm_simps fun_eq_iff) | 
| 1062 | 253 | |
| 254 | ||
| 255 | section {* Permutation Types *}
 | |
| 256 | ||
| 257 | text {*
 | |
| 258 |   Infix syntax for @{text permute} has higher precedence than
 | |
| 259 | addition, but lower than unary minus. | |
| 260 | *} | |
| 261 | ||
| 262 | class pt = | |
| 263 |   fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75)
 | |
| 264 | assumes permute_zero [simp]: "0 \<bullet> x = x" | |
| 265 | assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)" | |
| 266 | begin | |
| 267 | ||
| 268 | lemma permute_diff [simp]: | |
| 269 | shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x" | |
| 270 | unfolding diff_minus by simp | |
| 271 | ||
| 272 | lemma permute_minus_cancel [simp]: | |
| 273 | shows "p \<bullet> - p \<bullet> x = x" | |
| 274 | and "- p \<bullet> p \<bullet> x = x" | |
| 275 | unfolding permute_plus [symmetric] by simp_all | |
| 276 | ||
| 277 | lemma permute_swap_cancel [simp]: | |
| 278 | shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x" | |
| 279 | unfolding permute_plus [symmetric] | |
| 280 | by (simp add: swap_cancel) | |
| 281 | ||
| 282 | lemma permute_swap_cancel2 [simp]: | |
| 283 | shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x" | |
| 284 | unfolding permute_plus [symmetric] | |
| 285 | by (simp add: swap_commute) | |
| 286 | ||
| 287 | lemma inj_permute [simp]: | |
| 288 | shows "inj (permute p)" | |
| 289 | by (rule inj_on_inverseI) | |
| 290 | (rule permute_minus_cancel) | |
| 291 | ||
| 292 | lemma surj_permute [simp]: | |
| 293 | shows "surj (permute p)" | |
| 294 | by (rule surjI, rule permute_minus_cancel) | |
| 295 | ||
| 296 | lemma bij_permute [simp]: | |
| 297 | shows "bij (permute p)" | |
| 298 | by (rule bijI [OF inj_permute surj_permute]) | |
| 299 | ||
| 300 | lemma inv_permute: | |
| 301 | shows "inv (permute p) = permute (- p)" | |
| 302 | by (rule inv_equality) (simp_all) | |
| 303 | ||
| 304 | lemma permute_minus: | |
| 305 | shows "permute (- p) = inv (permute p)" | |
| 306 | by (simp add: inv_permute) | |
| 307 | ||
| 308 | lemma permute_eq_iff [simp]: | |
| 309 | shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y" | |
| 310 | by (rule inj_permute [THEN inj_eq]) | |
| 311 | ||
| 312 | end | |
| 313 | ||
| 314 | subsection {* Permutations for atoms *}
 | |
| 315 | ||
| 316 | instantiation atom :: pt | |
| 317 | begin | |
| 318 | ||
| 319 | definition | |
| 1879 | 320 | "p \<bullet> a = (Rep_perm p) a" | 
| 1062 | 321 | |
| 322 | instance | |
| 323 | apply(default) | |
| 324 | apply(simp_all add: permute_atom_def Rep_perm_simps) | |
| 325 | done | |
| 326 | ||
| 327 | end | |
| 328 | ||
| 329 | lemma sort_of_permute [simp]: | |
| 330 | shows "sort_of (p \<bullet> a) = sort_of a" | |
| 331 | unfolding permute_atom_def by (rule sort_of_Rep_perm) | |
| 332 | ||
| 333 | lemma swap_atom: | |
| 334 | shows "(a \<rightleftharpoons> b) \<bullet> c = | |
| 335 | (if sort_of a = sort_of b | |
| 336 | then (if c = a then b else if c = b then a else c) else c)" | |
| 337 | unfolding permute_atom_def | |
| 338 | by (simp add: Rep_perm_swap) | |
| 339 | ||
| 340 | lemma swap_atom_simps [simp]: | |
| 341 | "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b" | |
| 342 | "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a" | |
| 343 | "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c" | |
| 344 | unfolding swap_atom by simp_all | |
| 345 | ||
| 346 | lemma expand_perm_eq: | |
| 347 | fixes p q :: "perm" | |
| 348 | shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)" | |
| 349 | unfolding permute_atom_def | |
| 350 | by (metis Rep_perm_ext ext) | |
| 351 | ||
| 352 | ||
| 353 | subsection {* Permutations for permutations *}
 | |
| 354 | ||
| 355 | instantiation perm :: pt | |
| 356 | begin | |
| 357 | ||
| 358 | definition | |
| 359 | "p \<bullet> q = p + q - p" | |
| 360 | ||
| 361 | instance | |
| 362 | apply default | |
| 363 | apply (simp add: permute_perm_def) | |
| 364 | apply (simp add: permute_perm_def diff_minus minus_add add_assoc) | |
| 365 | done | |
| 366 | ||
| 367 | end | |
| 368 | ||
| 1879 | 369 | lemma permute_self: | 
| 370 | shows "p \<bullet> p = p" | |
| 371 | unfolding permute_perm_def | |
| 372 | by (simp add: diff_minus add_assoc) | |
| 1062 | 373 | |
| 374 | lemma permute_eqvt: | |
| 375 | shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)" | |
| 376 | unfolding permute_perm_def by simp | |
| 377 | ||
| 378 | lemma zero_perm_eqvt: | |
| 379 | shows "p \<bullet> (0::perm) = 0" | |
| 380 | unfolding permute_perm_def by simp | |
| 381 | ||
| 382 | lemma add_perm_eqvt: | |
| 383 | fixes p p1 p2 :: perm | |
| 384 | shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2" | |
| 385 | unfolding permute_perm_def | |
| 386 | by (simp add: expand_perm_eq) | |
| 387 | ||
| 388 | lemma swap_eqvt: | |
| 389 | shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)" | |
| 390 | unfolding permute_perm_def | |
| 391 | by (auto simp add: swap_atom expand_perm_eq) | |
| 392 | ||
| 2310 | 393 | lemma uminus_eqvt: | 
| 394 | fixes p q::"perm" | |
| 395 | shows "p \<bullet> (- q) = - (p \<bullet> q)" | |
| 396 | unfolding permute_perm_def | |
| 397 | by (simp add: diff_minus minus_add add_assoc) | |
| 1062 | 398 | |
| 399 | subsection {* Permutations for functions *}
 | |
| 400 | ||
| 401 | instantiation "fun" :: (pt, pt) pt | |
| 402 | begin | |
| 403 | ||
| 404 | definition | |
| 405 | "p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))" | |
| 406 | ||
| 407 | instance | |
| 408 | apply default | |
| 409 | apply (simp add: permute_fun_def) | |
| 410 | apply (simp add: permute_fun_def minus_add) | |
| 411 | done | |
| 412 | ||
| 413 | end | |
| 414 | ||
| 415 | lemma permute_fun_app_eq: | |
| 416 | shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)" | |
| 1879 | 417 | unfolding permute_fun_def by simp | 
| 1062 | 418 | |
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 419 | |
| 1062 | 420 | subsection {* Permutations for booleans *}
 | 
| 421 | ||
| 422 | instantiation bool :: pt | |
| 423 | begin | |
| 424 | ||
| 425 | definition "p \<bullet> (b::bool) = b" | |
| 426 | ||
| 427 | instance | |
| 428 | apply(default) | |
| 429 | apply(simp_all add: permute_bool_def) | |
| 430 | done | |
| 431 | ||
| 432 | end | |
| 433 | ||
| 434 | lemma Not_eqvt: | |
| 435 | shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))" | |
| 436 | by (simp add: permute_bool_def) | |
| 437 | ||
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 438 | lemma conj_eqvt: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 439 | shows "p \<bullet> (A \<and> B) = ((p \<bullet> A) \<and> (p \<bullet> B))" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 440 | by (simp add: permute_bool_def) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 441 | |
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 442 | lemma imp_eqvt: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 443 | shows "p \<bullet> (A \<longrightarrow> B) = ((p \<bullet> A) \<longrightarrow> (p \<bullet> B))" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 444 | by (simp add: permute_bool_def) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 445 | |
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 446 | lemma ex_eqvt: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 447 | shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 448 | unfolding permute_fun_def permute_bool_def | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 449 | by (auto, rule_tac x="p \<bullet> x" in exI, simp) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 450 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 451 | lemma all_eqvt: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 452 | shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 453 | unfolding permute_fun_def permute_bool_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 454 | by (auto, drule_tac x="p \<bullet> x" in spec, simp) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 455 | |
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 456 | lemma ex1_eqvt: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 457 | shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 458 | unfolding Ex1_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 459 | apply(simp add: ex_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 460 | unfolding permute_fun_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 461 | apply(simp add: conj_eqvt all_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 462 | unfolding permute_fun_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 463 | apply(simp add: imp_eqvt permute_bool_def) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 464 | done | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 465 | |
| 1557 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 466 | lemma permute_boolE: | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 467 | fixes P::"bool" | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 468 | shows "p \<bullet> P \<Longrightarrow> P" | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 469 | by (simp add: permute_bool_def) | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 470 | |
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 471 | lemma permute_boolI: | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 472 | fixes P::"bool" | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 473 | shows "P \<Longrightarrow> p \<bullet> P" | 
| 
fee2389789ad
moved infinite_Un into mainstream Isabelle; moved permute_boolI/E lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
1305diff
changeset | 474 | by(simp add: permute_bool_def) | 
| 1062 | 475 | |
| 476 | subsection {* Permutations for sets *}
 | |
| 477 | ||
| 478 | lemma permute_set_eq: | |
| 479 | fixes x::"'a::pt" | |
| 480 | and p::"perm" | |
| 481 |   shows "(p \<bullet> X) = {p \<bullet> x | x. x \<in> X}"
 | |
| 1879 | 482 | unfolding permute_fun_def | 
| 483 | unfolding permute_bool_def | |
| 484 | apply(auto simp add: mem_def) | |
| 1062 | 485 | apply(rule_tac x="- p \<bullet> x" in exI) | 
| 486 | apply(simp) | |
| 487 | done | |
| 488 | ||
| 489 | lemma permute_set_eq_image: | |
| 490 | shows "p \<bullet> X = permute p ` X" | |
| 1879 | 491 | unfolding permute_set_eq by auto | 
| 1062 | 492 | |
| 493 | lemma permute_set_eq_vimage: | |
| 494 | shows "p \<bullet> X = permute (- p) -` X" | |
| 1879 | 495 | unfolding permute_fun_def permute_bool_def | 
| 496 | unfolding vimage_def Collect_def mem_def .. | |
| 1062 | 497 | |
| 2588 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 498 | lemma permute_finite [simp]: | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 499 | shows "finite (p \<bullet> X) = finite X" | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 500 | apply(simp add: permute_set_eq_image) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 501 | apply(rule iffI) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 502 | apply(drule finite_imageD) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 503 | using inj_permute[where p="p"] | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 504 | apply(simp add: inj_on_def) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 505 | apply(assumption) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 506 | apply(rule finite_imageI) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 507 | apply(assumption) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 508 | done | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 509 | |
| 1062 | 510 | lemma swap_set_not_in: | 
| 511 | assumes a: "a \<notin> S" "b \<notin> S" | |
| 512 | shows "(a \<rightleftharpoons> b) \<bullet> S = S" | |
| 1879 | 513 | unfolding permute_set_eq | 
| 514 | using a by (auto simp add: swap_atom) | |
| 1062 | 515 | |
| 516 | lemma swap_set_in: | |
| 517 | assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b" | |
| 518 | shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S" | |
| 1879 | 519 | unfolding permute_set_eq | 
| 520 | using a by (auto simp add: swap_atom) | |
| 1062 | 521 | |
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 522 | lemma swap_set_in_eq: | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 523 | assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 524 |   shows "(a \<rightleftharpoons> b) \<bullet> S = (S - {a}) \<union> {b}"
 | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 525 | unfolding permute_set_eq | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 526 | using a by (auto simp add: swap_atom) | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 527 | |
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 528 | lemma swap_set_both_in: | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 529 | assumes a: "a \<in> S" "b \<in> S" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 530 | shows "(a \<rightleftharpoons> b) \<bullet> S = S" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 531 | unfolding permute_set_eq | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 532 | using a by (auto simp add: swap_atom) | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 533 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 534 | lemma mem_permute_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 535 | shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 536 | unfolding mem_def permute_fun_def permute_bool_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 537 | by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 538 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 539 | lemma mem_eqvt: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 540 | shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 541 | unfolding mem_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 542 | by (simp add: permute_fun_app_eq) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 543 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 544 | lemma empty_eqvt: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 545 |   shows "p \<bullet> {} = {}"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 546 | unfolding empty_def Collect_def | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 547 | by (simp add: permute_fun_def permute_bool_def) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 548 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 549 | lemma insert_eqvt: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 550 | shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 551 | unfolding permute_set_eq_image image_insert .. | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 552 | |
| 2672 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 553 | lemma Collect_eqvt: | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 554 |   shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
 | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 555 | unfolding Collect_def permute_fun_def .. | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 556 | |
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 557 | lemma inter_eqvt: | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 558 | shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)" | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 559 | unfolding Int_def | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 560 | apply(simp add: Collect_eqvt) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 561 | apply(simp add: permute_fun_def) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 562 | apply(simp add: conj_eqvt) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 563 | apply(simp add: mem_eqvt) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 564 | apply(simp add: permute_fun_def) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 565 | done | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 566 | |
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 567 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 568 | |
| 1062 | 569 | subsection {* Permutations for units *}
 | 
| 570 | ||
| 571 | instantiation unit :: pt | |
| 572 | begin | |
| 573 | ||
| 574 | definition "p \<bullet> (u::unit) = u" | |
| 575 | ||
| 1879 | 576 | instance | 
| 577 | by (default) (simp_all add: permute_unit_def) | |
| 1062 | 578 | |
| 579 | end | |
| 580 | ||
| 581 | ||
| 582 | subsection {* Permutations for products *}
 | |
| 583 | ||
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 584 | instantiation prod :: (pt, pt) pt | 
| 1062 | 585 | begin | 
| 586 | ||
| 587 | primrec | |
| 588 | permute_prod | |
| 589 | where | |
| 590 | Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)" | |
| 591 | ||
| 592 | instance | |
| 593 | by default auto | |
| 594 | ||
| 595 | end | |
| 596 | ||
| 597 | subsection {* Permutations for sums *}
 | |
| 598 | ||
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 599 | instantiation sum :: (pt, pt) pt | 
| 1062 | 600 | begin | 
| 601 | ||
| 602 | primrec | |
| 603 | permute_sum | |
| 604 | where | |
| 605 | "p \<bullet> (Inl x) = Inl (p \<bullet> x)" | |
| 606 | | "p \<bullet> (Inr y) = Inr (p \<bullet> y)" | |
| 607 | ||
| 1879 | 608 | instance | 
| 609 | by (default) (case_tac [!] x, simp_all) | |
| 1062 | 610 | |
| 611 | end | |
| 612 | ||
| 613 | subsection {* Permutations for lists *}
 | |
| 614 | ||
| 615 | instantiation list :: (pt) pt | |
| 616 | begin | |
| 617 | ||
| 618 | primrec | |
| 619 | permute_list | |
| 620 | where | |
| 621 | "p \<bullet> [] = []" | |
| 622 | | "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs" | |
| 623 | ||
| 1879 | 624 | instance | 
| 625 | by (default) (induct_tac [!] x, simp_all) | |
| 1062 | 626 | |
| 627 | end | |
| 628 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 629 | lemma set_eqvt: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 630 | shows "p \<bullet> (set xs) = set (p \<bullet> xs)" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 631 | by (induct xs) (simp_all add: empty_eqvt insert_eqvt) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 632 | |
| 1062 | 633 | subsection {* Permutations for options *}
 | 
| 634 | ||
| 635 | instantiation option :: (pt) pt | |
| 636 | begin | |
| 637 | ||
| 638 | primrec | |
| 639 | permute_option | |
| 640 | where | |
| 641 | "p \<bullet> None = None" | |
| 642 | | "p \<bullet> (Some x) = Some (p \<bullet> x)" | |
| 643 | ||
| 1879 | 644 | instance | 
| 645 | by (default) (induct_tac [!] x, simp_all) | |
| 1062 | 646 | |
| 647 | end | |
| 648 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 649 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 650 | subsection {* Permutations for fsets *}
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 651 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 652 | lemma permute_fset_rsp[quot_respect]: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 653 | shows "(op = ===> list_eq ===> list_eq) permute permute" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 654 | unfolding fun_rel_def | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 655 | by (simp add: set_eqvt[symmetric]) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 656 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 657 | instantiation fset :: (pt) pt | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 658 | begin | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 659 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 660 | quotient_definition | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 661 | "permute_fset :: perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 662 | is | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 663 | "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 664 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 665 | instance | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 666 | proof | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 667 | fix x :: "'a fset" and p q :: "perm" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 668 | show "0 \<bullet> x = x" by (descending) (simp) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 669 | show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" by (descending) (simp) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 670 | qed | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 671 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 672 | end | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 673 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 674 | lemma permute_fset [simp]: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 675 |   fixes S::"('a::pt) fset"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 676 |   shows "(p \<bullet> {||}) = ({||} ::('a::pt) fset)"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 677 | and "(p \<bullet> insert_fset x S) = insert_fset (p \<bullet> x) (p \<bullet> S)" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 678 | by (lifting permute_list.simps) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 679 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 680 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 681 | |
| 1062 | 682 | subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *}
 | 
| 683 | ||
| 684 | instantiation char :: pt | |
| 685 | begin | |
| 686 | ||
| 687 | definition "p \<bullet> (c::char) = c" | |
| 688 | ||
| 1879 | 689 | instance | 
| 690 | by (default) (simp_all add: permute_char_def) | |
| 1062 | 691 | |
| 692 | end | |
| 693 | ||
| 694 | instantiation nat :: pt | |
| 695 | begin | |
| 696 | ||
| 697 | definition "p \<bullet> (n::nat) = n" | |
| 698 | ||
| 1879 | 699 | instance | 
| 700 | by (default) (simp_all add: permute_nat_def) | |
| 1062 | 701 | |
| 702 | end | |
| 703 | ||
| 704 | instantiation int :: pt | |
| 705 | begin | |
| 706 | ||
| 707 | definition "p \<bullet> (i::int) = i" | |
| 708 | ||
| 1879 | 709 | instance | 
| 710 | by (default) (simp_all add: permute_int_def) | |
| 1062 | 711 | |
| 712 | end | |
| 713 | ||
| 714 | ||
| 715 | section {* Pure types *}
 | |
| 716 | ||
| 717 | text {* Pure types will have always empty support. *}
 | |
| 718 | ||
| 719 | class pure = pt + | |
| 720 | assumes permute_pure: "p \<bullet> x = x" | |
| 721 | ||
| 722 | text {* Types @{typ unit} and @{typ bool} are pure. *}
 | |
| 723 | ||
| 724 | instance unit :: pure | |
| 725 | proof qed (rule permute_unit_def) | |
| 726 | ||
| 727 | instance bool :: pure | |
| 728 | proof qed (rule permute_bool_def) | |
| 729 | ||
| 2635 
64b4cb2c2bf8
simple cases for string rule inductions
 Christian Urban <urbanc@in.tum.de> parents: 
2632diff
changeset | 730 | |
| 1062 | 731 | text {* Other type constructors preserve purity. *}
 | 
| 732 | ||
| 733 | instance "fun" :: (pure, pure) pure | |
| 734 | by default (simp add: permute_fun_def permute_pure) | |
| 735 | ||
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 736 | instance prod :: (pure, pure) pure | 
| 1062 | 737 | by default (induct_tac x, simp add: permute_pure) | 
| 738 | ||
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 739 | instance sum :: (pure, pure) pure | 
| 1062 | 740 | by default (induct_tac x, simp_all add: permute_pure) | 
| 741 | ||
| 742 | instance list :: (pure) pure | |
| 743 | by default (induct_tac x, simp_all add: permute_pure) | |
| 744 | ||
| 745 | instance option :: (pure) pure | |
| 746 | by default (induct_tac x, simp_all add: permute_pure) | |
| 747 | ||
| 748 | ||
| 749 | subsection {* Types @{typ char}, @{typ nat}, and @{typ int} *}
 | |
| 750 | ||
| 751 | instance char :: pure | |
| 752 | proof qed (rule permute_char_def) | |
| 753 | ||
| 754 | instance nat :: pure | |
| 755 | proof qed (rule permute_nat_def) | |
| 756 | ||
| 757 | instance int :: pure | |
| 758 | proof qed (rule permute_int_def) | |
| 759 | ||
| 760 | subsection {* Supp, Freshness and Supports *}
 | |
| 761 | ||
| 762 | context pt | |
| 763 | begin | |
| 764 | ||
| 765 | definition | |
| 766 | supp :: "'a \<Rightarrow> atom set" | |
| 767 | where | |
| 768 |   "supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}"
 | |
| 769 | ||
| 770 | end | |
| 771 | ||
| 772 | definition | |
| 773 |   fresh :: "atom \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55)
 | |
| 774 | where | |
| 775 | "a \<sharp> x \<equiv> a \<notin> supp x" | |
| 776 | ||
| 777 | lemma supp_conv_fresh: | |
| 778 |   shows "supp x = {a. \<not> a \<sharp> x}"
 | |
| 779 | unfolding fresh_def by simp | |
| 780 | ||
| 781 | lemma swap_rel_trans: | |
| 782 | assumes "sort_of a = sort_of b" | |
| 783 | assumes "sort_of b = sort_of c" | |
| 784 | assumes "(a \<rightleftharpoons> c) \<bullet> x = x" | |
| 785 | assumes "(b \<rightleftharpoons> c) \<bullet> x = x" | |
| 786 | shows "(a \<rightleftharpoons> b) \<bullet> x = x" | |
| 787 | proof (cases) | |
| 788 | assume "a = b \<or> c = b" | |
| 789 | with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto | |
| 790 | next | |
| 791 | assume *: "\<not> (a = b \<or> c = b)" | |
| 792 | have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x" | |
| 793 | using assms by simp | |
| 794 | also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)" | |
| 795 | using assms * by (simp add: swap_triple) | |
| 796 | finally show "(a \<rightleftharpoons> b) \<bullet> x = x" . | |
| 797 | qed | |
| 798 | ||
| 799 | lemma swap_fresh_fresh: | |
| 800 | assumes a: "a \<sharp> x" | |
| 801 | and b: "b \<sharp> x" | |
| 802 | shows "(a \<rightleftharpoons> b) \<bullet> x = x" | |
| 803 | proof (cases) | |
| 804 | assume asm: "sort_of a = sort_of b" | |
| 805 |   have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}" 
 | |
| 806 | using a b unfolding fresh_def supp_def by simp_all | |
| 807 |   then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp
 | |
| 808 | then obtain c | |
| 809 | where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b" | |
| 810 | by (rule obtain_atom) (auto) | |
| 811 | then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all) | |
| 812 | next | |
| 813 | assume "sort_of a \<noteq> sort_of b" | |
| 814 | then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp | |
| 815 | qed | |
| 816 | ||
| 817 | ||
| 818 | subsection {* supp and fresh are equivariant *}
 | |
| 819 | ||
| 820 | lemma finite_Collect_bij: | |
| 821 | assumes a: "bij f" | |
| 822 |   shows "finite {x. P (f x)} = finite {x. P x}"
 | |
| 823 | by (metis a finite_vimage_iff vimage_Collect_eq) | |
| 824 | ||
| 825 | lemma fresh_permute_iff: | |
| 826 | shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x" | |
| 827 | proof - | |
| 828 |   have "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
 | |
| 829 | unfolding fresh_def supp_def by simp | |
| 830 |   also have "\<dots> \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> p \<bullet> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
 | |
| 1879 | 831 | using bij_permute by (rule finite_Collect_bij[symmetric]) | 
| 1062 | 832 |   also have "\<dots> \<longleftrightarrow> finite {b. p \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> p \<bullet> x}"
 | 
| 833 | by (simp only: permute_eqvt [of p] swap_eqvt) | |
| 834 |   also have "\<dots> \<longleftrightarrow> finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"
 | |
| 835 | by (simp only: permute_eq_iff) | |
| 836 | also have "\<dots> \<longleftrightarrow> a \<sharp> x" | |
| 837 | unfolding fresh_def supp_def by simp | |
| 1879 | 838 | finally show "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x" . | 
| 1062 | 839 | qed | 
| 840 | ||
| 841 | lemma fresh_eqvt: | |
| 842 | shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)" | |
| 1879 | 843 | unfolding permute_bool_def | 
| 844 | by (simp add: fresh_permute_iff) | |
| 1062 | 845 | |
| 846 | lemma supp_eqvt: | |
| 847 | shows "p \<bullet> (supp x) = supp (p \<bullet> x)" | |
| 848 | unfolding supp_conv_fresh | |
| 1879 | 849 | unfolding Collect_def | 
| 850 | unfolding permute_fun_def | |
| 1062 | 851 | by (simp add: Not_eqvt fresh_eqvt) | 
| 852 | ||
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 853 | lemma fresh_permute_left: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 854 | shows "a \<sharp> p \<bullet> x \<longleftrightarrow> - p \<bullet> a \<sharp> x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 855 | proof | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 856 | assume "a \<sharp> p \<bullet> x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 857 | then have "- p \<bullet> a \<sharp> - p \<bullet> p \<bullet> x" by (simp only: fresh_permute_iff) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 858 | then show "- p \<bullet> a \<sharp> x" by simp | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 859 | next | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 860 | assume "- p \<bullet> a \<sharp> x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 861 | then have "p \<bullet> - p \<bullet> a \<sharp> p \<bullet> x" by (simp only: fresh_permute_iff) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 862 | then show "a \<sharp> p \<bullet> x" by simp | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 863 | qed | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 864 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 865 | |
| 1062 | 866 | subsection {* supports *}
 | 
| 867 | ||
| 868 | definition | |
| 869 | supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80) | |
| 870 | where | |
| 871 | "S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)" | |
| 872 | ||
| 873 | lemma supp_is_subset: | |
| 874 | fixes S :: "atom set" | |
| 875 | and x :: "'a::pt" | |
| 876 | assumes a1: "S supports x" | |
| 877 | and a2: "finite S" | |
| 878 | shows "(supp x) \<subseteq> S" | |
| 879 | proof (rule ccontr) | |
| 1879 | 880 | assume "\<not> (supp x \<subseteq> S)" | 
| 1062 | 881 | then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto | 
| 1879 | 882 | from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" unfolding supports_def by auto | 
| 883 |   then have "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto
 | |
| 1062 | 884 |   with a2 have "finite {b. (a \<rightleftharpoons> b)\<bullet>x \<noteq> x}" by (simp add: finite_subset)
 | 
| 885 | then have "a \<notin> (supp x)" unfolding supp_def by simp | |
| 886 | with b1 show False by simp | |
| 887 | qed | |
| 888 | ||
| 889 | lemma supports_finite: | |
| 890 | fixes S :: "atom set" | |
| 891 | and x :: "'a::pt" | |
| 892 | assumes a1: "S supports x" | |
| 893 | and a2: "finite S" | |
| 894 | shows "finite (supp x)" | |
| 895 | proof - | |
| 896 | have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset) | |
| 897 | then show "finite (supp x)" using a2 by (simp add: finite_subset) | |
| 898 | qed | |
| 899 | ||
| 900 | lemma supp_supports: | |
| 901 | fixes x :: "'a::pt" | |
| 902 | shows "(supp x) supports x" | |
| 1879 | 903 | unfolding supports_def | 
| 904 | proof (intro strip) | |
| 1062 | 905 | fix a b | 
| 906 | assume "a \<notin> (supp x) \<and> b \<notin> (supp x)" | |
| 907 | then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def) | |
| 1879 | 908 | then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh) | 
| 1062 | 909 | qed | 
| 910 | ||
| 911 | lemma supp_is_least_supports: | |
| 912 | fixes S :: "atom set" | |
| 913 | and x :: "'a::pt" | |
| 914 | assumes a1: "S supports x" | |
| 915 | and a2: "finite S" | |
| 916 | and a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'" | |
| 917 | shows "(supp x) = S" | |
| 918 | proof (rule equalityI) | |
| 919 | show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset) | |
| 920 | with a2 have fin: "finite (supp x)" by (rule rev_finite_subset) | |
| 921 | have "(supp x) supports x" by (rule supp_supports) | |
| 922 | with fin a3 show "S \<subseteq> supp x" by blast | |
| 923 | qed | |
| 924 | ||
| 925 | lemma subsetCI: | |
| 926 | shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B" | |
| 927 | by auto | |
| 928 | ||
| 929 | lemma finite_supp_unique: | |
| 930 | assumes a1: "S supports x" | |
| 931 | assumes a2: "finite S" | |
| 932 | assumes a3: "\<And>a b. \<lbrakk>a \<in> S; b \<notin> S; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x \<noteq> x" | |
| 933 | shows "(supp x) = S" | |
| 934 | using a1 a2 | |
| 935 | proof (rule supp_is_least_supports) | |
| 936 | fix S' | |
| 937 | assume "finite S'" and "S' supports x" | |
| 938 | show "S \<subseteq> S'" | |
| 939 | proof (rule subsetCI) | |
| 940 | fix a | |
| 941 | assume "a \<in> S" and "a \<notin> S'" | |
| 942 | have "finite (S \<union> S')" | |
| 943 | using `finite S` `finite S'` by simp | |
| 944 | then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a" | |
| 945 | by (rule obtain_atom) | |
| 946 | then have "b \<notin> S" and "b \<notin> S'" and "sort_of a = sort_of b" | |
| 947 | by simp_all | |
| 948 | then have "(a \<rightleftharpoons> b) \<bullet> x = x" | |
| 949 | using `a \<notin> S'` `S' supports x` by (simp add: supports_def) | |
| 950 | moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x" | |
| 951 | using `a \<in> S` `b \<notin> S` `sort_of a = sort_of b` | |
| 952 | by (rule a3) | |
| 953 | ultimately show "False" by simp | |
| 954 | qed | |
| 955 | qed | |
| 956 | ||
| 2475 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 957 | section {* Support w.r.t. relations *}
 | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 958 | |
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 959 | text {* 
 | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 960 | This definition is used for unquotient types, where | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 961 | alpha-equivalence does not coincide with equality. | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 962 | *} | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 963 | |
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 964 | definition | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 965 |   "supp_rel R x = {a. infinite {b. \<not>(R ((a \<rightleftharpoons> b) \<bullet> x) x)}}"
 | 
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 966 | |
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 967 | |
| 
486d4647bb37
supp-proofs work except for CoreHaskell and Modules (induct is probably not finding the correct instance)
 Christian Urban <urbanc@in.tum.de> parents: 
2470diff
changeset | 968 | |
| 1062 | 969 | section {* Finitely-supported types *}
 | 
| 970 | ||
| 971 | class fs = pt + | |
| 972 | assumes finite_supp: "finite (supp x)" | |
| 973 | ||
| 974 | lemma pure_supp: | |
| 975 |   shows "supp (x::'a::pure) = {}"
 | |
| 976 | unfolding supp_def by (simp add: permute_pure) | |
| 977 | ||
| 978 | lemma pure_fresh: | |
| 979 | fixes x::"'a::pure" | |
| 980 | shows "a \<sharp> x" | |
| 981 | unfolding fresh_def by (simp add: pure_supp) | |
| 982 | ||
| 983 | instance pure < fs | |
| 984 | by default (simp add: pure_supp) | |
| 985 | ||
| 986 | ||
| 987 | subsection  {* Type @{typ atom} is finitely-supported. *}
 | |
| 988 | ||
| 989 | lemma supp_atom: | |
| 990 |   shows "supp a = {a}"
 | |
| 991 | apply (rule finite_supp_unique) | |
| 992 | apply (clarsimp simp add: supports_def) | |
| 993 | apply simp | |
| 994 | apply simp | |
| 995 | done | |
| 996 | ||
| 997 | lemma fresh_atom: | |
| 998 | shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b" | |
| 999 | unfolding fresh_def supp_atom by simp | |
| 1000 | ||
| 1001 | instance atom :: fs | |
| 1002 | by default (simp add: supp_atom) | |
| 1003 | ||
| 1933 
9eab1dfc14d2
moved lemmas from FSet.thy to do with atom to Nominal2_Base, and to do with 'a::at set to Nominal2_Atoms; moved Nominal2_Eqvt.thy one up to be loaded before Nominal2_Atoms
 Christian Urban <urbanc@in.tum.de> parents: 
1932diff
changeset | 1004 | |
| 1062 | 1005 | section {* Type @{typ perm} is finitely-supported. *}
 | 
| 1006 | ||
| 1007 | lemma perm_swap_eq: | |
| 1008 | shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)" | |
| 1009 | unfolding permute_perm_def | |
| 1010 | by (metis add_diff_cancel minus_perm_def) | |
| 1011 | ||
| 1012 | lemma supports_perm: | |
| 1013 |   shows "{a. p \<bullet> a \<noteq> a} supports p"
 | |
| 1014 | unfolding supports_def | |
| 1879 | 1015 | unfolding perm_swap_eq | 
| 1016 | by (simp add: swap_eqvt) | |
| 1062 | 1017 | |
| 1018 | lemma finite_perm_lemma: | |
| 1019 |   shows "finite {a::atom. p \<bullet> a \<noteq> a}"
 | |
| 1020 | using finite_Rep_perm [of p] | |
| 1021 | unfolding permute_atom_def . | |
| 1022 | ||
| 1023 | lemma supp_perm: | |
| 1024 |   shows "supp p = {a. p \<bullet> a \<noteq> a}"
 | |
| 1025 | apply (rule finite_supp_unique) | |
| 1026 | apply (rule supports_perm) | |
| 1027 | apply (rule finite_perm_lemma) | |
| 1028 | apply (simp add: perm_swap_eq swap_eqvt) | |
| 1029 | apply (auto simp add: expand_perm_eq swap_atom) | |
| 1030 | done | |
| 1031 | ||
| 1032 | lemma fresh_perm: | |
| 1033 | shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a" | |
| 1879 | 1034 | unfolding fresh_def | 
| 1035 | by (simp add: supp_perm) | |
| 1062 | 1036 | |
| 1037 | lemma supp_swap: | |
| 1038 |   shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})"
 | |
| 1039 | by (auto simp add: supp_perm swap_atom) | |
| 1040 | ||
| 1041 | lemma fresh_zero_perm: | |
| 1042 | shows "a \<sharp> (0::perm)" | |
| 1043 | unfolding fresh_perm by simp | |
| 1044 | ||
| 1045 | lemma supp_zero_perm: | |
| 1046 |   shows "supp (0::perm) = {}"
 | |
| 1047 | unfolding supp_perm by simp | |
| 1048 | ||
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1049 | lemma fresh_plus_perm: | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1050 | fixes p q::perm | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1051 | assumes "a \<sharp> p" "a \<sharp> q" | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1052 | shows "a \<sharp> (p + q)" | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1053 | using assms | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1054 | unfolding fresh_def | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1055 | by (auto simp add: supp_perm) | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1056 | |
| 1062 | 1057 | lemma supp_plus_perm: | 
| 1058 | fixes p q::perm | |
| 1059 | shows "supp (p + q) \<subseteq> supp p \<union> supp q" | |
| 1060 | by (auto simp add: supp_perm) | |
| 1061 | ||
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1062 | lemma fresh_minus_perm: | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1063 | fixes p::perm | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1064 | shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p" | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1065 | unfolding fresh_def | 
| 1879 | 1066 | unfolding supp_perm | 
| 1067 | apply(simp) | |
| 1068 | apply(metis permute_minus_cancel) | |
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1069 | done | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1070 | |
| 1062 | 1071 | lemma supp_minus_perm: | 
| 1072 | fixes p::perm | |
| 1073 | shows "supp (- p) = supp p" | |
| 1087 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1074 | unfolding supp_conv_fresh | 
| 
bb7f4457091a
moved some lemmas to Nominal; updated all files
 Christian Urban <urbanc@in.tum.de> parents: 
1062diff
changeset | 1075 | by (simp add: fresh_minus_perm) | 
| 1062 | 1076 | |
| 1077 | instance perm :: fs | |
| 1078 | by default (simp add: supp_perm finite_perm_lemma) | |
| 1079 | ||
| 1305 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1080 | lemma plus_perm_eq: | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1081 | fixes p q::"perm" | 
| 1879 | 1082 |   assumes asm: "supp p \<inter> supp q = {}"
 | 
| 1305 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1083 | shows "p + q = q + p" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1084 | unfolding expand_perm_eq | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1085 | proof | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1086 | fix a::"atom" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1087 | show "(p + q) \<bullet> a = (q + p) \<bullet> a" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1088 | proof - | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1089 |     { assume "a \<notin> supp p" "a \<notin> supp q"
 | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1090 | then have "(p + q) \<bullet> a = (q + p) \<bullet> a" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1091 | by (simp add: supp_perm) | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1092 | } | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1093 | moreover | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1094 |     { assume a: "a \<in> supp p" "a \<notin> supp q"
 | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1095 | then have "p \<bullet> a \<in> supp p" by (simp add: supp_perm) | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1096 | then have "p \<bullet> a \<notin> supp q" using asm by auto | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1097 | with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1098 | by (simp add: supp_perm) | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1099 | } | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1100 | moreover | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1101 |     { assume a: "a \<notin> supp p" "a \<in> supp q"
 | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1102 | then have "q \<bullet> a \<in> supp q" by (simp add: supp_perm) | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1103 | then have "q \<bullet> a \<notin> supp p" using asm by auto | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1104 | with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1105 | by (simp add: supp_perm) | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1106 | } | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1107 | ultimately show "(p + q) \<bullet> a = (q + p) \<bullet> a" | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1108 | using asm by blast | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1109 | qed | 
| 
61319a9af976
updated (added lemma about commuting permutations)
 Christian Urban <urbanc@in.tum.de> parents: 
1258diff
changeset | 1110 | qed | 
| 1062 | 1111 | |
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1112 | lemma supp_plus_perm_eq: | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1113 | fixes p q::perm | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1114 |   assumes asm: "supp p \<inter> supp q = {}"
 | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1115 | shows "supp (p + q) = supp p \<union> supp q" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1116 | proof - | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1117 |   { fix a::"atom"
 | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1118 | assume "a \<in> supp p" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1119 | then have "a \<notin> supp q" using asm by auto | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1120 | then have "a \<in> supp (p + q)" using `a \<in> supp p` | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1121 | by (simp add: supp_perm) | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1122 | } | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1123 | moreover | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1124 |   { fix a::"atom"
 | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1125 | assume "a \<in> supp q" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1126 | then have "a \<notin> supp p" using asm by auto | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1127 | then have "a \<in> supp (q + p)" using `a \<in> supp q` | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1128 | by (simp add: supp_perm) | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1129 | then have "a \<in> supp (p + q)" using asm plus_perm_eq | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1130 | by metis | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1131 | } | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1132 | ultimately have "supp p \<union> supp q \<subseteq> supp (p + q)" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1133 | by blast | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1134 | then show "supp (p + q) = supp p \<union> supp q" using supp_plus_perm | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1135 | by blast | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1136 | qed | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1137 | |
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1138 | |
| 1062 | 1139 | section {* Finite Support instances for other types *}
 | 
| 1140 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1141 | |
| 1062 | 1142 | subsection {* Type @{typ "'a \<times> 'b"} is finitely-supported. *}
 | 
| 1143 | ||
| 1144 | lemma supp_Pair: | |
| 1145 | shows "supp (x, y) = supp x \<union> supp y" | |
| 1146 | by (simp add: supp_def Collect_imp_eq Collect_neg_eq) | |
| 1147 | ||
| 1148 | lemma fresh_Pair: | |
| 1149 | shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y" | |
| 1150 | by (simp add: fresh_def supp_Pair) | |
| 1151 | ||
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1152 | lemma supp_Unit: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1153 |   shows "supp () = {}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1154 | by (simp add: supp_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1155 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1156 | lemma fresh_Unit: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1157 | shows "a \<sharp> ()" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1158 | by (simp add: fresh_def supp_Unit) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1159 | |
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 1160 | instance prod :: (fs, fs) fs | 
| 1062 | 1161 | apply default | 
| 1162 | apply (induct_tac x) | |
| 1163 | apply (simp add: supp_Pair finite_supp) | |
| 1164 | done | |
| 1165 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1166 | |
| 1062 | 1167 | subsection {* Type @{typ "'a + 'b"} is finitely supported *}
 | 
| 1168 | ||
| 1169 | lemma supp_Inl: | |
| 1170 | shows "supp (Inl x) = supp x" | |
| 1171 | by (simp add: supp_def) | |
| 1172 | ||
| 1173 | lemma supp_Inr: | |
| 1174 | shows "supp (Inr x) = supp x" | |
| 1175 | by (simp add: supp_def) | |
| 1176 | ||
| 1177 | lemma fresh_Inl: | |
| 1178 | shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x" | |
| 1179 | by (simp add: fresh_def supp_Inl) | |
| 1180 | ||
| 1181 | lemma fresh_Inr: | |
| 1182 | shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y" | |
| 1183 | by (simp add: fresh_def supp_Inr) | |
| 1184 | ||
| 2378 
2f13fe48c877
updated to new Isabelle; made FSet more "quiet"
 Christian Urban <urbanc@in.tum.de> parents: 
2310diff
changeset | 1185 | instance sum :: (fs, fs) fs | 
| 1062 | 1186 | apply default | 
| 1187 | apply (induct_tac x) | |
| 1188 | apply (simp_all add: supp_Inl supp_Inr finite_supp) | |
| 1189 | done | |
| 1190 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1191 | |
| 1062 | 1192 | subsection {* Type @{typ "'a option"} is finitely supported *}
 | 
| 1193 | ||
| 1194 | lemma supp_None: | |
| 1195 |   shows "supp None = {}"
 | |
| 1196 | by (simp add: supp_def) | |
| 1197 | ||
| 1198 | lemma supp_Some: | |
| 1199 | shows "supp (Some x) = supp x" | |
| 1200 | by (simp add: supp_def) | |
| 1201 | ||
| 1202 | lemma fresh_None: | |
| 1203 | shows "a \<sharp> None" | |
| 1204 | by (simp add: fresh_def supp_None) | |
| 1205 | ||
| 1206 | lemma fresh_Some: | |
| 1207 | shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x" | |
| 1208 | by (simp add: fresh_def supp_Some) | |
| 1209 | ||
| 1210 | instance option :: (fs) fs | |
| 1211 | apply default | |
| 1212 | apply (induct_tac x) | |
| 1213 | apply (simp_all add: supp_None supp_Some finite_supp) | |
| 1214 | done | |
| 1215 | ||
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1216 | |
| 1062 | 1217 | subsubsection {* Type @{typ "'a list"} is finitely supported *}
 | 
| 1218 | ||
| 1219 | lemma supp_Nil: | |
| 1220 |   shows "supp [] = {}"
 | |
| 1221 | by (simp add: supp_def) | |
| 1222 | ||
| 1223 | lemma supp_Cons: | |
| 1224 | shows "supp (x # xs) = supp x \<union> supp xs" | |
| 1225 | by (simp add: supp_def Collect_imp_eq Collect_neg_eq) | |
| 1226 | ||
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1227 | lemma supp_append: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1228 | shows "supp (xs @ ys) = supp xs \<union> supp ys" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1229 | by (induct xs) (auto simp add: supp_Nil supp_Cons) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1230 | |
| 1062 | 1231 | lemma fresh_Nil: | 
| 1232 | shows "a \<sharp> []" | |
| 1233 | by (simp add: fresh_def supp_Nil) | |
| 1234 | ||
| 1235 | lemma fresh_Cons: | |
| 1236 | shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs" | |
| 1237 | by (simp add: fresh_def supp_Cons) | |
| 1238 | ||
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1239 | lemma fresh_append: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1240 | shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1241 | by (induct xs) (simp_all add: fresh_Nil fresh_Cons) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1242 | |
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1243 | |
| 1062 | 1244 | instance list :: (fs) fs | 
| 1245 | apply default | |
| 1246 | apply (induct_tac x) | |
| 1247 | apply (simp_all add: supp_Nil supp_Cons finite_supp) | |
| 1248 | done | |
| 1249 | ||
| 2588 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1250 | lemma supp_of_atom_list: | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1251 | fixes as::"atom list" | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1252 | shows "supp as = set as" | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1253 | by (induct as) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1254 | (simp_all add: supp_Nil supp_Cons supp_atom) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1255 | |
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1256 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1257 | section {* Support and Freshness for Applications *}
 | 
| 1062 | 1258 | |
| 1879 | 1259 | lemma fresh_conv_MOST: | 
| 1260 | shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)" | |
| 1261 | unfolding fresh_def supp_def | |
| 1262 | unfolding MOST_iff_cofinite by simp | |
| 1263 | ||
| 2003 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1264 | lemma supp_subset_fresh: | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1265 | assumes a: "\<And>a. a \<sharp> x \<Longrightarrow> a \<sharp> y" | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1266 | shows "supp y \<subseteq> supp x" | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1267 | using a | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1268 | unfolding fresh_def | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1269 | by blast | 
| 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1270 | |
| 1879 | 1271 | lemma fresh_fun_app: | 
| 1272 | assumes "a \<sharp> f" and "a \<sharp> x" | |
| 1273 | shows "a \<sharp> f x" | |
| 2003 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1274 | using assms | 
| 1879 | 1275 | unfolding fresh_conv_MOST | 
| 1276 | unfolding permute_fun_app_eq | |
| 1277 | by (elim MOST_rev_mp, simp) | |
| 1278 | ||
| 1062 | 1279 | lemma supp_fun_app: | 
| 1280 | shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)" | |
| 1879 | 1281 | using fresh_fun_app | 
| 1282 | unfolding fresh_def | |
| 1283 | by auto | |
| 1284 | ||
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1285 | text {* Equivariant Functions *}
 | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1286 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1287 | definition | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1288 | "eqvt f \<equiv> \<forall>p. p \<bullet> f = f" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1289 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1290 | lemma eqvtI: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1291 | shows "(\<And>p. p \<bullet> f \<equiv> f) \<Longrightarrow> eqvt f" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1292 | unfolding eqvt_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1293 | by simp | 
| 2003 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1294 | |
| 1941 | 1295 | lemma supp_fun_eqvt: | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1296 | assumes a: "eqvt f" | 
| 1941 | 1297 |   shows "supp f = {}"
 | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1298 | using a | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1299 | unfolding eqvt_def | 
| 1941 | 1300 | unfolding supp_def | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1301 | by simp | 
| 1941 | 1302 | |
| 1062 | 1303 | lemma fresh_fun_eqvt_app: | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1304 | assumes a: "eqvt f" | 
| 1062 | 1305 | shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x" | 
| 1306 | proof - | |
| 1941 | 1307 |   from a have "supp f = {}" by (simp add: supp_fun_eqvt)
 | 
| 1879 | 1308 | then show "a \<sharp> x \<Longrightarrow> a \<sharp> f x" | 
| 1062 | 1309 | unfolding fresh_def | 
| 2003 
b53e98bfb298
added lemmas establishing the support of finite sets of finitely supported elements
 Christian Urban <urbanc@in.tum.de> parents: 
1973diff
changeset | 1310 | using supp_fun_app by auto | 
| 1062 | 1311 | qed | 
| 1312 | ||
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1313 | text {* equivariance of a function at a given argument *}
 | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1314 | definition | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1315 | "eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1316 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1317 | lemma supp_eqvt_at: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1318 | assumes asm: "eqvt_at f x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1319 | and fin: "finite (supp x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1320 | shows "supp (f x) \<subseteq> supp x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1321 | apply(rule supp_is_subset) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1322 | unfolding supports_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1323 | unfolding fresh_def[symmetric] | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1324 | using asm | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1325 | apply(simp add: eqvt_at_def) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1326 | apply(simp add: swap_fresh_fresh) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1327 | apply(rule fin) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1328 | done | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1329 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1330 | lemma finite_supp_eqvt_at: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1331 | assumes asm: "eqvt_at f x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1332 | and fin: "finite (supp x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1333 | shows "finite (supp (f x))" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1334 | apply(rule finite_subset) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1335 | apply(rule supp_eqvt_at[OF asm fin]) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1336 | apply(rule fin) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1337 | done | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1338 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1339 | lemma fresh_eqvt_at: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1340 | assumes asm: "eqvt_at f x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1341 | and fin: "finite (supp x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1342 | and fresh: "a \<sharp> x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1343 | shows "a \<sharp> f x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1344 | using fresh | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1345 | unfolding fresh_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1346 | using supp_eqvt_at[OF asm fin] | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1347 | by auto | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1348 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1349 | text {* helper functions for nominal_functions *}
 | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1350 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1351 | lemma the_default_eqvt: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1352 | assumes unique: "\<exists>!x. P x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1353 | shows "(p \<bullet> (THE_default d P)) = (THE_default d (p \<bullet> P))" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1354 | apply(rule THE_default1_equality [symmetric]) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1355 | apply(rule_tac p="-p" in permute_boolE) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1356 | apply(simp add: ex1_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1357 | apply(rule unique) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1358 | apply(rule_tac p="-p" in permute_boolE) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1359 | apply(rule subst[OF permute_fun_app_eq]) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1360 | apply(simp) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1361 | apply(rule THE_defaultI'[OF unique]) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1362 | done | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1363 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1364 | lemma fundef_ex1_eqvt: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1365 | fixes x::"'a::pt" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1366 | assumes f_def: "f == (\<lambda>x::'a. THE_default d (G x))" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1367 | assumes eqvt: "eqvt G" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1368 | assumes ex1: "\<exists>!y. G x y" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1369 | shows "(p \<bullet> (f x)) = f (p \<bullet> x)" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1370 | apply(simp only: f_def) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1371 | apply(subst the_default_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1372 | apply(rule ex1) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1373 | using eqvt | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1374 | unfolding eqvt_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1375 | apply(simp add: permute_fun_app_eq) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1376 | done | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1377 | |
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1378 | lemma fundef_ex1_eqvt_at: | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1379 | fixes x::"'a::pt" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1380 | assumes f_def: "f == (\<lambda>x::'a. THE_default d (G x))" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1381 | assumes eqvt: "eqvt G" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1382 | assumes ex1: "\<exists>!y. G x y" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1383 | shows "eqvt_at f x" | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1384 | unfolding eqvt_at_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1385 | using assms | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1386 | by (auto intro: fundef_ex1_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1387 | |
| 1962 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 1388 | |
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1389 | section {* Support of Finite Sets of Finitely Supported Elements *}
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1390 | |
| 2657 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1391 | text {* support and freshness for atom sets *}
 | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1392 | |
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1393 | lemma supp_finite_atom_set: | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1394 | fixes S::"atom set" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1395 | assumes "finite S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1396 | shows "supp S = S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1397 | apply(rule finite_supp_unique) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1398 | apply(simp add: supports_def) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1399 | apply(simp add: swap_set_not_in) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1400 | apply(rule assms) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1401 | apply(simp add: swap_set_in) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1402 | done | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1403 | |
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1404 | lemma fresh_finite_atom_set: | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1405 | fixes S::"atom set" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1406 | assumes "finite S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1407 | shows "a \<sharp> S \<longleftrightarrow> a \<notin> S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1408 | unfolding fresh_def | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1409 | by (simp add: supp_finite_atom_set[OF assms]) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 1410 | |
| 2679 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1411 | lemma fresh_minus_atom_set: | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1412 | fixes S::"atom set" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1413 | assumes "finite S" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1414 | shows "a \<sharp> S - T \<longleftrightarrow> (a \<notin> T \<longrightarrow> a \<sharp> S)" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1415 | unfolding fresh_def | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1416 | by (auto simp add: supp_finite_atom_set assms) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1417 | |
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1418 | |
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1419 | lemma Union_fresh: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1420 | shows "a \<sharp> S \<Longrightarrow> a \<sharp> (\<Union>x \<in> S. supp x)" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1421 | unfolding Union_image_eq[symmetric] | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1422 | apply(rule_tac f="\<lambda>S. \<Union> supp ` S" in fresh_fun_eqvt_app) | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1423 | unfolding eqvt_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1424 | unfolding permute_fun_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1425 | apply(simp) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1426 | unfolding UNION_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1427 | unfolding Collect_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1428 | unfolding Bex_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1429 | apply(simp add: ex_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1430 | unfolding permute_fun_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1431 | apply(simp add: conj_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1432 | apply(simp add: mem_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1433 | apply(simp add: supp_eqvt) | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1434 | unfolding permute_fun_def | 
| 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1435 | apply(simp) | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1436 | apply(assumption) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1437 | done | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1438 | |
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1439 | lemma Union_supports_set: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1440 | shows "(\<Union>x \<in> S. supp x) supports S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1441 | proof - | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1442 |   { fix a b
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1443 | have "\<forall>x \<in> S. (a \<rightleftharpoons> b) \<bullet> x = x \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> S = S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1444 | unfolding permute_set_eq by force | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1445 | } | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1446 | then show "(\<Union>x \<in> S. supp x) supports S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1447 | unfolding supports_def | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1448 | by (simp add: fresh_def[symmetric] swap_fresh_fresh) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1449 | qed | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1450 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1451 | lemma Union_of_finite_supp_sets: | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1452 |   fixes S::"('a::fs set)"
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1453 | assumes fin: "finite S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1454 | shows "finite (\<Union>x\<in>S. supp x)" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1455 | using fin by (induct) (auto simp add: finite_supp) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1456 | |
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1457 | lemma Union_included_in_supp: | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1458 |   fixes S::"('a::fs set)"
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1459 | assumes fin: "finite S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1460 | shows "(\<Union>x\<in>S. supp x) \<subseteq> supp S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1461 | proof - | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1462 | have "(\<Union>x\<in>S. supp x) = supp (\<Union>x\<in>S. supp x)" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1463 | by (rule supp_finite_atom_set[symmetric]) | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1464 | (rule Union_of_finite_supp_sets[OF fin]) | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1465 | also have "\<dots> \<subseteq> supp S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1466 | by (rule supp_subset_fresh) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1467 | (simp add: Union_fresh) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1468 | finally show "(\<Union>x\<in>S. supp x) \<subseteq> supp S" . | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1469 | qed | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1470 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1471 | lemma supp_of_finite_sets: | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1472 |   fixes S::"('a::fs set)"
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1473 | assumes fin: "finite S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1474 | shows "(supp S) = (\<Union>x\<in>S. supp x)" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1475 | apply(rule subset_antisym) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1476 | apply(rule supp_is_subset) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1477 | apply(rule Union_supports_set) | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1478 | apply(rule Union_of_finite_supp_sets[OF fin]) | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1479 | apply(rule Union_included_in_supp[OF fin]) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1480 | done | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1481 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1482 | lemma finite_sets_supp: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1483 |   fixes S::"('a::fs set)"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1484 | assumes "finite S" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1485 | shows "finite (supp S)" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1486 | using assms | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1487 | by (simp only: supp_of_finite_sets Union_of_finite_supp_sets) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1488 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1489 | lemma supp_of_finite_union: | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1490 |   fixes S T::"('a::fs) set"
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1491 | assumes fin1: "finite S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1492 | and fin2: "finite T" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1493 | shows "supp (S \<union> T) = supp S \<union> supp T" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1494 | using fin1 fin2 | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1495 | by (simp add: supp_of_finite_sets) | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1496 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1497 | lemma supp_of_finite_insert: | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1498 |   fixes S::"('a::fs) set"
 | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1499 | assumes fin: "finite S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1500 | shows "supp (insert x S) = supp x \<union> supp S" | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1501 | using fin | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1502 | by (simp add: supp_of_finite_sets) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1503 | |
| 2588 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1504 | lemma fresh_finite_insert: | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1505 |   fixes S::"('a::fs) set"
 | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1506 | assumes fin: "finite S" | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1507 | shows "a \<sharp> (insert x S) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S" | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1508 | using fin unfolding fresh_def | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1509 | by (simp add: supp_of_finite_insert) | 
| 
8f5420681039
completed the strong exhausts rules for Foo2 using general lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2587diff
changeset | 1510 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1511 | lemma supp_set_empty: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1512 |   shows "supp {} = {}"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1513 | unfolding supp_def | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1514 | by (simp add: empty_eqvt) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1515 | |
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1516 | lemma fresh_set_empty: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1517 |   shows "a \<sharp> {}"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1518 | by (simp add: fresh_def supp_set_empty) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1519 | |
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1520 | lemma supp_set: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1521 |   fixes xs :: "('a::fs) list"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1522 | shows "supp (set xs) = supp xs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1523 | apply(induct xs) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1524 | apply(simp add: supp_set_empty supp_Nil) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1525 | apply(simp add: supp_Cons supp_of_finite_insert) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1526 | done | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1527 | |
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1528 | lemma fresh_set: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1529 |   fixes xs :: "('a::fs) list"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1530 | shows "a \<sharp> (set xs) \<longleftrightarrow> a \<sharp> xs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1531 | unfolding fresh_def | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1532 | by (simp add: supp_set) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1533 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1534 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1535 | subsection {* Type @{typ "'a fset"} is finitely supported *}
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1536 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1537 | lemma fset_eqvt: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1538 | shows "p \<bullet> (fset S) = fset (p \<bullet> S)" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1539 | by (lifting set_eqvt) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1540 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1541 | lemma supp_fset [simp]: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1542 | shows "supp (fset S) = supp S" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1543 | unfolding supp_def | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1544 | by (simp add: fset_eqvt fset_cong) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1545 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1546 | lemma supp_empty_fset [simp]: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1547 |   shows "supp {||} = {}"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1548 | unfolding supp_def | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1549 | by simp | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1550 | |
| 2641 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1551 | lemma fresh_empty_fset: | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1552 |   shows "a \<sharp> {||}"
 | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1553 | unfolding fresh_def | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1554 | by (simp) | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1555 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1556 | lemma supp_insert_fset [simp]: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1557 | fixes x::"'a::fs" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1558 | and S::"'a fset" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1559 | shows "supp (insert_fset x S) = supp x \<union> supp S" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1560 | apply(subst supp_fset[symmetric]) | 
| 2587 
78623a0f294b
tuned proof to reduce number of warnings
 Christian Urban <urbanc@in.tum.de> parents: 
2586diff
changeset | 1561 | apply(simp add: supp_of_finite_insert) | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1562 | done | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1563 | |
| 2641 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1564 | lemma fresh_insert_fset: | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1565 | fixes x::"'a::fs" | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1566 | and S::"'a fset" | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1567 | shows "a \<sharp> insert_fset x S \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S" | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1568 | unfolding fresh_def | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1569 | by (simp) | 
| 
592d17e26e09
some further lemmas for fsets
 Christian Urban <urbanc@in.tum.de> parents: 
2635diff
changeset | 1570 | |
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1571 | lemma fset_finite_supp: | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1572 |   fixes S::"('a::fs) fset"
 | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1573 | shows "finite (supp S)" | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1574 | by (induct S) (simp_all add: finite_supp) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1575 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1576 | |
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1577 | instance fset :: (fs) fs | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1578 | apply (default) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1579 | apply (rule fset_finite_supp) | 
| 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 1580 | done | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1581 | |
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 1582 | |
| 2632 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1583 | section {* Freshness and Fresh-Star *}
 | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1584 | |
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1585 | lemma fresh_Unit_elim: | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1586 | shows "(a \<sharp> () \<Longrightarrow> PROP C) \<equiv> PROP C" | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1587 | by (simp add: fresh_Unit) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1588 | |
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1589 | lemma fresh_Pair_elim: | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1590 | shows "(a \<sharp> (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> a \<sharp> y \<Longrightarrow> PROP C)" | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1591 | by rule (simp_all add: fresh_Pair) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1592 | |
| 2632 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1593 | (* this rule needs to be added before the fresh_prodD is *) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1594 | (* added to the simplifier with mksimps *) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1595 | lemma [simp]: | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1596 | shows "a \<sharp> x1 \<Longrightarrow> a \<sharp> x2 \<Longrightarrow> a \<sharp> (x1, x2)" | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1597 | by (simp add: fresh_Pair) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1598 | |
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1599 | lemma fresh_PairD: | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1600 | shows "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> x" | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1601 | and "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> y" | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1602 | by (simp_all add: fresh_Pair) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1603 | |
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1604 | ML {*
 | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1605 |   val mksimps_pairs = (@{const_name Nominal2_Base.fresh}, @{thms fresh_PairD}) :: mksimps_pairs;
 | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1606 | *} | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1607 | |
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1608 | declaration {* fn _ =>
 | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1609 | Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) | 
| 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1610 | *} | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1611 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1612 | text {* The fresh-star generalisation of fresh is used in strong
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1613 | induction principles. *} | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1614 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1615 | definition | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1616 |   fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80)
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1617 | where | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1618 | "as \<sharp>* x \<equiv> \<forall>a \<in> as. a \<sharp> x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1619 | |
| 2507 | 1620 | lemma fresh_star_supp_conv: | 
| 1621 | shows "supp x \<sharp>* y \<Longrightarrow> supp y \<sharp>* x" | |
| 1622 | by (auto simp add: fresh_star_def fresh_def) | |
| 1623 | ||
| 2675 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1624 | lemma fresh_star_perm_set_conv: | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1625 | fixes p::"perm" | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1626 | assumes fresh: "as \<sharp>* p" | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1627 | and fin: "finite as" | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1628 | shows "supp p \<sharp>* as" | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1629 | apply(rule fresh_star_supp_conv) | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1630 | apply(simp add: supp_finite_atom_set fin fresh) | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1631 | done | 
| 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1632 | |
| 2679 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1633 | lemma fresh_star_atom_set_conv: | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1634 | fixes p::"perm" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1635 | assumes fresh: "as \<sharp>* bs" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1636 | and fin: "finite as" "finite bs" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1637 | shows "bs \<sharp>* as" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1638 | using fresh | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1639 | unfolding fresh_star_def fresh_def | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1640 | by (auto simp add: supp_finite_atom_set fin) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1641 | |
| 2675 
68ccf847507d
defined properly substitution
 Christian Urban <urbanc@in.tum.de> parents: 
2672diff
changeset | 1642 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1643 | lemma fresh_star_Pair: | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1644 | shows "as \<sharp>* (x, y) = (as \<sharp>* x \<and> as \<sharp>* y)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1645 | by (auto simp add: fresh_star_def fresh_Pair) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1646 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1647 | lemma fresh_star_list: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1648 | shows "as \<sharp>* (xs @ ys) \<longleftrightarrow> as \<sharp>* xs \<and> as \<sharp>* ys" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1649 | and "as \<sharp>* (x # xs) \<longleftrightarrow> as \<sharp>* x \<and> as \<sharp>* xs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1650 | and "as \<sharp>* []" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1651 | by (auto simp add: fresh_star_def fresh_Nil fresh_Cons fresh_append) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1652 | |
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1653 | lemma fresh_star_set: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1654 |   fixes xs::"('a::fs) list"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1655 | shows "as \<sharp>* set xs \<longleftrightarrow> as \<sharp>* xs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1656 | unfolding fresh_star_def | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1657 | by (simp add: fresh_set) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1658 | |
| 2611 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1659 | lemma fresh_star_singleton: | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1660 | fixes a::"atom" | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1661 |   shows "as \<sharp>* {a} \<longleftrightarrow> as \<sharp>* a"
 | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1662 | by (simp add: fresh_star_def fresh_finite_insert fresh_set_empty) | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1663 | |
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1664 | lemma fresh_star_fset: | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1665 |   fixes xs::"('a::fs) list"
 | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1666 | shows "as \<sharp>* fset S \<longleftrightarrow> as \<sharp>* S" | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1667 | by (simp add: fresh_star_def fresh_def) | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 1668 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1669 | lemma fresh_star_Un: | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1670 | shows "(as \<union> bs) \<sharp>* x = (as \<sharp>* x \<and> bs \<sharp>* x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1671 | by (auto simp add: fresh_star_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1672 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1673 | lemma fresh_star_insert: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1674 | shows "(insert a as) \<sharp>* x = (a \<sharp> x \<and> as \<sharp>* x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1675 | by (auto simp add: fresh_star_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1676 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1677 | lemma fresh_star_Un_elim: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1678 | "((as \<union> bs) \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (as \<sharp>* x \<Longrightarrow> bs \<sharp>* x \<Longrightarrow> PROP C)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1679 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1680 | apply(rule) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1681 | apply(erule meta_mp) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1682 | apply(auto) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1683 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1684 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1685 | lemma fresh_star_insert_elim: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1686 | "(insert a as \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> as \<sharp>* x \<Longrightarrow> PROP C)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1687 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1688 | by rule (simp_all add: fresh_star_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1689 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1690 | lemma fresh_star_empty_elim: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1691 |   "({} \<sharp>* x \<Longrightarrow> PROP C) \<equiv> PROP C"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1692 | by (simp add: fresh_star_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1693 | |
| 2632 
e8732350a29f
added small example for strong inductions; functions still need a sorry
 Christian Urban <urbanc@in.tum.de> parents: 
2614diff
changeset | 1694 | lemma fresh_star_Unit_elim: | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1695 | shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1696 | by (simp add: fresh_star_def fresh_Unit) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1697 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1698 | lemma fresh_star_Pair_elim: | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1699 | shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)" | 
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1700 | by (rule, simp_all add: fresh_star_Pair) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1701 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1702 | lemma fresh_star_zero: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1703 | shows "as \<sharp>* (0::perm)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1704 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1705 | by (simp add: fresh_zero_perm) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1706 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1707 | lemma fresh_star_plus: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1708 | fixes p q::perm | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1709 | shows "\<lbrakk>a \<sharp>* p; a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1710 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1711 | by (simp add: fresh_plus_perm) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1712 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1713 | lemma fresh_star_permute_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1714 | shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1715 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1716 | by (metis mem_permute_iff permute_minus_cancel(1) fresh_permute_iff) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1717 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1718 | lemma fresh_star_eqvt: | 
| 2663 
54aade5d0fe6
moved high level code from LamTest into the main libraries.
 Christian Urban <urbanc@in.tum.de> parents: 
2659diff
changeset | 1719 | shows "p \<bullet> (as \<sharp>* x) \<longleftrightarrow> (p \<bullet> as) \<sharp>* (p \<bullet> x)" | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1720 | unfolding fresh_star_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1721 | unfolding Ball_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1722 | apply(simp add: all_eqvt) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1723 | apply(subst permute_fun_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1724 | apply(simp add: imp_eqvt fresh_eqvt mem_eqvt) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1725 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1726 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1727 | lemma at_fresh_star_inter: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1728 | assumes a: "(p \<bullet> bs) \<sharp>* bs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1729 | and b: "finite bs" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1730 |   shows "p \<bullet> bs \<inter> bs = {}"
 | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1731 | using a b | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1732 | unfolding fresh_star_def | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1733 | unfolding fresh_def | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1734 | by (auto simp add: supp_finite_atom_set) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1735 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1736 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1737 | section {* Induction principle for permutations *}
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1738 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1739 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1740 | lemma perm_struct_induct[consumes 1, case_names zero swap]: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1741 | assumes S: "supp p \<subseteq> S" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1742 | and zero: "P 0" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1743 | and swap: "\<And>p a b. \<lbrakk>P p; supp p \<subseteq> S; a \<in> S; b \<in> S; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1744 | shows "P p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1745 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1746 | have "finite (supp p)" by (simp add: finite_supp) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1747 | then show "P p" using S | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1748 | proof(induct A\<equiv>"supp p" arbitrary: p rule: finite_psubset_induct) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1749 | case (psubset p) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1750 | then have ih: "\<And>q. supp q \<subset> supp p \<Longrightarrow> P q" by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1751 | have as: "supp p \<subseteq> S" by fact | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1752 |     { assume "supp p = {}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1753 | then have "p = 0" by (simp add: supp_perm expand_perm_eq) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1754 | then have "P p" using zero by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1755 | } | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1756 | moreover | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1757 |     { assume "supp p \<noteq> {}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1758 | then obtain a where a0: "a \<in> supp p" by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1759 | then have a1: "p \<bullet> a \<in> S" "a \<in> S" "sort_of (p \<bullet> a) = sort_of a" "p \<bullet> a \<noteq> a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1760 | using as by (auto simp add: supp_atom supp_perm swap_atom) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1761 | let ?q = "(p \<bullet> a \<rightleftharpoons> a) + p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1762 | have a2: "supp ?q \<subseteq> supp p" unfolding supp_perm by (auto simp add: swap_atom) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1763 | moreover | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1764 | have "a \<notin> supp ?q" by (simp add: supp_perm) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1765 | then have "supp ?q \<noteq> supp p" using a0 by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1766 | ultimately have "supp ?q \<subset> supp p" using a2 by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1767 | then have "P ?q" using ih by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1768 | moreover | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1769 | have "supp ?q \<subseteq> S" using as a2 by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1770 | ultimately have "P ((p \<bullet> a \<rightleftharpoons> a) + ?q)" using as a1 swap by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1771 | moreover | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1772 | have "p = (p \<bullet> a \<rightleftharpoons> a) + ?q" by (simp add: expand_perm_eq) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1773 | ultimately have "P p" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1774 | } | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1775 | ultimately show "P p" by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1776 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1777 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1778 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1779 | lemma perm_simple_struct_induct[case_names zero swap]: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1780 | assumes zero: "P 0" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1781 | and swap: "\<And>p a b. \<lbrakk>P p; a \<noteq> b; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> P ((a \<rightleftharpoons> b) + p)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1782 | shows "P p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1783 | by (rule_tac S="supp p" in perm_struct_induct) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1784 | (auto intro: zero swap) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1785 | |
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1786 | lemma perm_struct_induct2[consumes 1, case_names zero swap plus]: | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1787 | assumes S: "supp p \<subseteq> S" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1788 | assumes zero: "P 0" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1789 | assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b; a \<in> S; b \<in> S\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1790 | assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2; supp p1 \<subseteq> S; supp p2 \<subseteq> S\<rbrakk> \<Longrightarrow> P (p1 + p2)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1791 | shows "P p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1792 | using S | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1793 | by (induct p rule: perm_struct_induct) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1794 | (auto intro: zero plus swap simp add: supp_swap) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1795 | |
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1796 | lemma perm_simple_struct_induct2[case_names zero swap plus]: | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1797 | assumes zero: "P 0" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1798 | assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1799 | assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1800 | shows "P p" | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1801 | by (rule_tac S="supp p" in perm_struct_induct2) | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1802 | (auto intro: zero swap plus) | 
| 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1803 | |
| 2679 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1804 | lemma supp_perm_singleton: | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1805 | fixes p::"perm" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1806 |   shows "supp p \<subseteq> {b} \<longleftrightarrow> p = 0"
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1807 | proof - | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1808 |   { assume "supp p \<subseteq> {b}"
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1809 | then have "p = 0" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1810 | by (induct p rule: perm_struct_induct) (simp_all) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1811 | } | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1812 |   then show "supp p \<subseteq> {b} \<longleftrightarrow> p = 0" by (auto simp add: supp_zero_perm)
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1813 | qed | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1814 | |
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1815 | lemma supp_perm_pair: | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1816 | fixes p::"perm" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1817 |   shows "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)"
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1818 | proof - | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1819 |   { assume "supp p \<subseteq> {a, b}"
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1820 | then have "p = 0 \<or> p = (b \<rightleftharpoons> a)" | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1821 | apply (induct p rule: perm_struct_induct) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1822 | apply (auto simp add: swap_cancel supp_zero_perm supp_swap) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1823 | apply (simp add: swap_commute) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1824 | done | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1825 | } | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1826 |   then show "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)" 
 | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1827 | by (auto simp add: supp_zero_perm supp_swap split: if_splits) | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1828 | qed | 
| 
e003e5e36bae
added Minimal file to test things
 Christian Urban <urbanc@in.tum.de> parents: 
2675diff
changeset | 1829 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1830 | lemma supp_perm_eq: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1831 | assumes "(supp x) \<sharp>* p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1832 | shows "p \<bullet> x = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1833 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1834 |   from assms have "supp p \<subseteq> {a. a \<sharp> x}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1835 | unfolding supp_perm fresh_star_def fresh_def by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1836 | then show "p \<bullet> x = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1837 | proof (induct p rule: perm_struct_induct) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1838 | case zero | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1839 | show "0 \<bullet> x = x" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1840 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1841 | case (swap p a b) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1842 | then have "a \<sharp> x" "b \<sharp> x" "p \<bullet> x = x" by simp_all | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1843 | then show "((a \<rightleftharpoons> b) + p) \<bullet> x = x" by (simp add: swap_fresh_fresh) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1844 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1845 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1846 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1847 | lemma supp_perm_eq_test: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1848 | assumes "(supp x) \<sharp>* p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1849 | shows "p \<bullet> x = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1850 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1851 |   from assms have "supp p \<subseteq> {a. a \<sharp> x}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1852 | unfolding supp_perm fresh_star_def fresh_def by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1853 | then show "p \<bullet> x = x" | 
| 2669 
1d1772a89026
the function translating lambda terms to locally nameless lambda terms; still needs a stronger abs_eq_iff lemma...at the moment only proved for restrictions
 Christian Urban <urbanc@in.tum.de> parents: 
2668diff
changeset | 1854 | proof (induct p rule: perm_struct_induct2) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1855 | case zero | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1856 | show "0 \<bullet> x = x" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1857 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1858 | case (swap a b) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1859 | then have "a \<sharp> x" "b \<sharp> x" by simp_all | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1860 | then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1861 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1862 | case (plus p1 p2) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1863 | have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+ | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1864 | then show "(p1 + p2) \<bullet> x = x" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1865 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1866 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1867 | |
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1868 | lemma perm_supp_eq: | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1869 | assumes a: "(supp p) \<sharp>* x" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1870 | shows "p \<bullet> x = x" | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1871 | by (rule supp_perm_eq) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1872 | (simp add: fresh_star_supp_conv a) | 
| 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1873 | |
| 2659 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1874 | lemma supp_perm_perm_eq: | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1875 | assumes a: "\<forall>a \<in> supp x. p \<bullet> a = q \<bullet> a" | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1876 | shows "p \<bullet> x = q \<bullet> x" | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1877 | proof - | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1878 | from a have "\<forall>a \<in> supp x. (-q + p) \<bullet> a = a" by simp | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1879 | then have "\<forall>a \<in> supp x. a \<notin> supp (-q + p)" | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1880 | unfolding supp_perm by simp | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1881 | then have "supp x \<sharp>* (-q + p)" | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1882 | unfolding fresh_star_def fresh_def by simp | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1883 | then have "(-q + p) \<bullet> x = x" by (simp only: supp_perm_eq) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1884 | then show "p \<bullet> x = q \<bullet> x" | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1885 | by (metis permute_minus_cancel permute_plus) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1886 | qed | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 1887 | |
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1888 | lemma atom_set_perm_eq: | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1889 | assumes a: "as \<sharp>* p" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1890 | shows "p \<bullet> as = as" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1891 | proof - | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1892 |   from a have "supp p \<subseteq> {a. a \<notin> as}"
 | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1893 | unfolding supp_perm fresh_star_def fresh_def by auto | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1894 | then show "p \<bullet> as = as" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1895 | proof (induct p rule: perm_struct_induct) | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1896 | case zero | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1897 | show "0 \<bullet> as = as" by simp | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1898 | next | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1899 | case (swap p a b) | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1900 | then have "a \<notin> as" "b \<notin> as" "p \<bullet> as = as" by simp_all | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1901 | then show "((a \<rightleftharpoons> b) + p) \<bullet> as = as" by (simp add: swap_set_not_in) | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1902 | qed | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 1903 | qed | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1904 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1905 | section {* Avoiding of atom sets *}
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1906 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1907 | text {* 
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1908 | For every set of atoms, there is another set of atoms | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1909 | avoiding a finitely supported c and there is a permutation | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1910 | which 'translates' between both sets. | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1911 | *} | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1912 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1913 | lemma at_set_avoiding_aux: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1914 | fixes Xs::"atom set" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1915 | and As::"atom set" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1916 | assumes b: "Xs \<subseteq> As" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1917 | and c: "finite As" | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1918 |   shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) = (Xs \<union> (p \<bullet> Xs))"
 | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1919 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1920 | from b c have "finite Xs" by (rule finite_subset) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1921 | then show ?thesis using b | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1922 | proof (induct rule: finite_subset_induct) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1923 | case empty | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1924 |     have "0 \<bullet> {} \<inter> As = {}" by simp
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1925 | moreover | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1926 |     have "supp (0::perm) = {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm)
 | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1927 | ultimately show ?case by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1928 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1929 | case (insert x Xs) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1930 | then obtain p where | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1931 |       p1: "(p \<bullet> Xs) \<inter> As = {}" and 
 | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1932 | p2: "supp p = (Xs \<union> (p \<bullet> Xs))" by blast | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1933 | from `x \<in> As` p1 have "x \<notin> p \<bullet> Xs" by fast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1934 | with `x \<notin> Xs` p2 have "x \<notin> supp p" by fast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1935 | hence px: "p \<bullet> x = x" unfolding supp_perm by simp | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1936 | have "finite (As \<union> p \<bullet> Xs \<union> supp p)" | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1937 | using `finite As` `finite Xs` | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1938 | by (simp add: permute_set_eq_image finite_supp) | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1939 | then obtain y where "y \<notin> (As \<union> p \<bullet> Xs \<union> supp p)" "sort_of y = sort_of x" | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1940 | by (rule obtain_atom) | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1941 | hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "y \<notin> supp p" "sort_of y = sort_of x" | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1942 | by simp_all | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1943 | hence py: "p \<bullet> y = y" "x \<noteq> y" using `x \<in> As` | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1944 | by (auto simp add: supp_perm) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1945 | let ?q = "(x \<rightleftharpoons> y) + p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1946 | have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1947 | unfolding insert_eqvt | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1948 | using `p \<bullet> x = x` `sort_of y = sort_of x` | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1949 | using `x \<notin> p \<bullet> Xs` `y \<notin> p \<bullet> Xs` | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1950 | by (simp add: swap_atom swap_set_not_in) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1951 |     have "?q \<bullet> insert x Xs \<inter> As = {}"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1952 |       using `y \<notin> As` `p \<bullet> Xs \<inter> As = {}`
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1953 | unfolding q by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1954 | moreover | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1955 |     have "supp (x \<rightleftharpoons> y) \<inter> supp p = {}" using px py `sort_of y = sort_of x`
 | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1956 | unfolding supp_swap by (simp add: supp_perm) | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1957 | then have "supp ?q = (supp (x \<rightleftharpoons> y) \<union> supp p)" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1958 | by (simp add: supp_plus_perm_eq) | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1959 | then have "supp ?q = insert x Xs \<union> ?q \<bullet> insert x Xs" | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1960 | using p2 `sort_of y = sort_of x` `x \<noteq> y` unfolding q supp_swap | 
| 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1961 | by auto | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1962 | ultimately show ?case by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1963 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1964 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1965 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1966 | lemma at_set_avoiding: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1967 | assumes a: "finite Xs" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1968 | and b: "finite (supp c)" | 
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 1969 | obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) = (Xs \<union> (p \<bullet> Xs))" | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1970 | using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"] | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1971 | unfolding fresh_star_def fresh_def by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1972 | |
| 2589 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1973 | lemma at_set_avoiding1: | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1974 | assumes "finite xs" | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1975 | and "finite (supp c)" | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1976 | shows "\<exists>p. (p \<bullet> xs) \<sharp>* c" | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1977 | using assms | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1978 | apply(erule_tac c="c" in at_set_avoiding) | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1979 | apply(auto) | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1980 | done | 
| 
9781db0e2196
completed proofs in Foo2
 Christian Urban <urbanc@in.tum.de> parents: 
2588diff
changeset | 1981 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1982 | lemma at_set_avoiding2: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1983 | assumes "finite xs" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1984 | and "finite (supp c)" "finite (supp x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1985 | and "xs \<sharp>* x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1986 | shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1987 | using assms | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1988 | apply(erule_tac c="(c, x)" in at_set_avoiding) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1989 | apply(simp add: supp_Pair) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1990 | apply(rule_tac x="p" in exI) | 
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 1991 | apply(simp add: fresh_star_Pair) | 
| 2507 | 1992 | apply(rule fresh_star_supp_conv) | 
| 1993 | apply(auto simp add: fresh_star_def) | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1994 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 1995 | |
| 2573 | 1996 | lemma at_set_avoiding3: | 
| 1997 | assumes "finite xs" | |
| 1998 | and "finite (supp c)" "finite (supp x)" | |
| 1999 | and "xs \<sharp>* x" | |
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 2000 | shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p \<and> supp p = xs \<union> (p \<bullet> xs)" | 
| 2586 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2001 | using assms | 
| 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2002 | apply(erule_tac c="(c, x)" in at_set_avoiding) | 
| 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2003 | apply(simp add: supp_Pair) | 
| 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2004 | apply(rule_tac x="p" in exI) | 
| 2591 
35c570891a3a
isarfied some of the high-level proofs
 Christian Urban <urbanc@in.tum.de> parents: 
2589diff
changeset | 2005 | apply(simp add: fresh_star_Pair) | 
| 2586 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2006 | apply(rule fresh_star_supp_conv) | 
| 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2007 | apply(auto simp add: fresh_star_def) | 
| 
3ebc7ecfb0dd
disabled the Foo examples, because of heavy work
 Christian Urban <urbanc@in.tum.de> parents: 
2573diff
changeset | 2008 | done | 
| 2573 | 2009 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2010 | lemma at_set_avoiding2_atom: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2011 | assumes "finite (supp c)" "finite (supp x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2012 | and b: "a \<sharp> x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2013 | shows "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2014 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2015 |   have a: "{a} \<sharp>* x" unfolding fresh_star_def by (simp add: b)
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2016 |   obtain p where p1: "(p \<bullet> {a}) \<sharp>* c" and p2: "supp x \<sharp>* p"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2017 |     using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2018 | have c: "(p \<bullet> a) \<sharp> c" using p1 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2019 | unfolding fresh_star_def Ball_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2020 | by(erule_tac x="p \<bullet> a" in allE) (simp add: permute_set_eq) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2021 | hence "p \<bullet> a \<sharp> c \<and> supp x \<sharp>* p" using p2 by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2022 | then show "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" by blast | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2023 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2024 | |
| 2614 
0d7a1703fe28
a stronger statement for at_set_avoiding
 Christian Urban <urbanc@in.tum.de> parents: 
2611diff
changeset | 2025 | |
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2026 | section {* Renaming permutations *}
 | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2027 | |
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2028 | lemma set_renaming_perm: | 
| 2659 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2029 | assumes b: "finite bs" | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2030 | shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" | 
| 2659 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2031 | using b | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2032 | proof (induct) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2033 | case empty | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2034 |   have "(\<forall>b \<in> {}. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> {} \<union> p \<bullet> {}"
 | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2035 | by (simp add: permute_set_eq supp_perm) | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2036 |   then show "\<exists>q. (\<forall>b \<in> {}. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> {} \<union> p \<bullet> {}" by blast
 | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2037 | next | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2038 | case (insert a bs) | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2039 | then have " \<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> p \<bullet> bs" by simp | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2040 | then obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> p \<bullet> bs" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2041 | by (metis empty_subsetI insert(3) supp_swap) | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2042 |   { assume 1: "q \<bullet> a = p \<bullet> a"
 | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2043 | have "\<forall>b \<in> (insert a bs). q \<bullet> b = p \<bullet> b" using 1 * by simp | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2044 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2045 | have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2046 | using ** by (auto simp add: insert_eqvt) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2047 | ultimately | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2048 | have "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2049 | } | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2050 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2051 |   { assume 2: "q \<bullet> a \<noteq> p \<bullet> a"
 | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2052 | def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q" | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2053 | have "\<forall>b \<in> insert a bs. q' \<bullet> b = p \<bullet> b" using 2 * `a \<notin> bs` unfolding q'_def | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2054 | by (auto simp add: swap_atom) | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2055 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2056 |     { have "{q \<bullet> a, p \<bullet> a} \<subseteq> insert a bs \<union> p \<bullet> insert a bs"
 | 
| 2659 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2057 | using ** | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2058 | apply (auto simp add: supp_perm insert_eqvt) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2059 | apply (subgoal_tac "q \<bullet> a \<in> bs \<union> p \<bullet> bs") | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2060 | apply(auto)[1] | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2061 | 	apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}")
 | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2062 | apply(blast) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2063 | apply(simp) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2064 | done | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2065 | then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by (simp add: supp_swap) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2066 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2067 | have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2068 | using ** by (auto simp add: insert_eqvt) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2069 | ultimately | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2070 | have "supp q' \<subseteq> insert a bs \<union> p \<bullet> insert a bs" | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2071 | unfolding q'_def using supp_plus_perm by blast | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2072 | } | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2073 | ultimately | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2074 | have "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" by blast | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2075 | } | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2076 | ultimately show "\<exists>q. (\<forall>b \<in> insert a bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2077 | by blast | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2078 | qed | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2079 | |
| 2672 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2080 | lemma set_renaming_perm2: | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2081 | shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2082 | proof - | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2083 | have "finite (bs \<inter> supp p)" by (simp add: finite_supp) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2084 | then obtain q | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2085 | where *: "\<forall>b \<in> bs \<inter> supp p. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> (bs \<inter> supp p) \<union> (p \<bullet> (bs \<inter> supp p))" | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2086 | using set_renaming_perm by blast | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2087 | from ** have "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by (auto simp add: inter_eqvt) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2088 | moreover | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2089 | have "\<forall>b \<in> bs - supp p. q \<bullet> b = p \<bullet> b" | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2090 | apply(auto) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2091 | apply(subgoal_tac "b \<notin> supp q") | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2092 | apply(simp add: fresh_def[symmetric]) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2093 | apply(simp add: fresh_perm) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2094 | apply(clarify) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2095 | apply(rotate_tac 2) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2096 | apply(drule subsetD[OF **]) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2097 | apply(simp add: inter_eqvt supp_eqvt permute_self) | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2098 | done | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2099 | ultimately have "(\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" using * by auto | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2100 | then show "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" by blast | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2101 | qed | 
| 
7e7662890477
removed finiteness assumption from set_rename_perm
 Christian Urban <urbanc@in.tum.de> parents: 
2669diff
changeset | 2102 | |
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2103 | lemma list_renaming_perm: | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2104 | shows "\<exists>q. (\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set bs \<union> (p \<bullet> set bs)" | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2105 | proof (induct bs) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2106 | case Nil | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2107 | have "(\<forall>b \<in> set []. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> set [] \<union> p \<bullet> set []" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2108 | by (simp add: supp_zero_perm) | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2109 | then show "\<exists>q. (\<forall>b \<in> set []. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set [] \<union> p \<bullet> (set [])" by blast | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2110 | next | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2111 | case (Cons a bs) | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2112 | then have " \<exists>q. (\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set bs \<union> p \<bullet> (set bs)" by simp | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2113 | then obtain q where *: "\<forall>b \<in> set bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> set bs \<union> p \<bullet> (set bs)" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2114 | by (blast) | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2115 |   { assume 1: "a \<in> set bs"
 | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2116 | have "q \<bullet> a = p \<bullet> a" using * 1 by (induct bs) (auto) | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2117 | then have "\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b" using * by simp | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2118 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2119 | have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" using ** by (auto simp add: insert_eqvt) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2120 | ultimately | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2121 | have "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2122 | } | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2123 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2124 |   { assume 2: "a \<notin> set bs"
 | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2125 | def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q" | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2126 | have "\<forall>b \<in> set (a # bs). q' \<bullet> b = p \<bullet> b" | 
| 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2127 | unfolding q'_def using 2 * `a \<notin> set bs` by (auto simp add: swap_atom) | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2128 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2129 |     { have "{q \<bullet> a, p \<bullet> a} \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))"
 | 
| 2659 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2130 | using ** | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2131 | apply (auto simp add: supp_perm insert_eqvt) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2132 | apply (subgoal_tac "q \<bullet> a \<in> set bs \<union> p \<bullet> set bs") | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2133 | apply(auto)[1] | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2134 | 	apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}")
 | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2135 | apply(blast) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2136 | apply(simp) | 
| 
619ecb57db38
strengthened renaming lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2657diff
changeset | 2137 | done | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2138 | then have "supp (q \<bullet> a \<rightleftharpoons> p \<bullet> a) \<subseteq> set (a # bs) \<union> p \<bullet> set (a # bs)" by (simp add: supp_swap) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2139 | moreover | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2140 | have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2141 | using ** by (auto simp add: insert_eqvt) | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2142 | ultimately | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2143 | have "supp q' \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2144 | unfolding q'_def using supp_plus_perm by blast | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2145 | } | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2146 | ultimately | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2147 | have "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" by blast | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2148 | } | 
| 2668 
92c001d93225
modified the renaming_perm lemmas
 Christian Urban <urbanc@in.tum.de> parents: 
2663diff
changeset | 2149 | ultimately show "\<exists>q. (\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" | 
| 2599 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2150 | by blast | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2151 | qed | 
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2152 | |
| 
d6fe94028a5d
moved general theorems into the libraries
 Christian Urban <urbanc@in.tum.de> parents: 
2591diff
changeset | 2153 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2154 | section {* Concrete Atoms Types *}
 | 
| 1962 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2155 | |
| 1972 | 2156 | text {*
 | 
| 2157 |   Class @{text at_base} allows types containing multiple sorts of atoms.
 | |
| 2158 |   Class @{text at} only allows types with a single sort.
 | |
| 2159 | *} | |
| 2160 | ||
| 1962 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2161 | class at_base = pt + | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2162 | fixes atom :: "'a \<Rightarrow> atom" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2163 | assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2164 | assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2165 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2166 | class at = at_base + | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2167 | assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2168 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2169 | lemma supp_at_base: | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2170 | fixes a::"'a::at_base" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2171 |   shows "supp a = {atom a}"
 | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2172 | by (simp add: supp_atom [symmetric] supp_def atom_eqvt) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2173 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2174 | lemma fresh_at_base: | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2175 | shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2176 | unfolding fresh_def by (simp add: supp_at_base) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2177 | |
| 2609 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2178 | lemma fresh_atom_at_base: | 
| 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2179 | fixes b::"'a::at_base" | 
| 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2180 | shows "a \<sharp> atom b \<longleftrightarrow> a \<sharp> b" | 
| 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2181 | by (simp add: fresh_def supp_at_base supp_atom) | 
| 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2182 | |
| 2611 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 2183 | lemma fresh_star_atom_at_base: | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 2184 | fixes b::"'a::at_base" | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 2185 | shows "as \<sharp>* atom b \<longleftrightarrow> as \<sharp>* b" | 
| 
3d101f2f817c
simple cases for strong inducts done; infrastructure for the difficult ones is there
 Christian Urban <urbanc@in.tum.de> parents: 
2609diff
changeset | 2186 | by (simp add: fresh_star_def fresh_atom_at_base) | 
| 2609 
666ffc8a92a9
freshness theorem in strong exhausts; (temporarily includes a cheat_tac to make all tests go through)
 Christian Urban <urbanc@in.tum.de> parents: 
2599diff
changeset | 2187 | |
| 1962 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2188 | instance at_base < fs | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2189 | proof qed (simp add: supp_at_base) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2190 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2191 | lemma at_base_infinite [simp]: | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2192 | shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U") | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2193 | proof | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2194 | obtain a :: 'a where "True" by auto | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2195 | assume "finite ?U" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2196 | hence "finite (atom ` ?U)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2197 | by (rule finite_imageI) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2198 | then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2199 | by (rule obtain_atom) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2200 | from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2201 | unfolding atom_eqvt [symmetric] | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2202 | by (simp add: swap_atom) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2203 | hence "b \<in> atom ` ?U" by simp | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2204 | with b(1) show "False" by simp | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2205 | qed | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2206 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2207 | lemma swap_at_base_simps [simp]: | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2208 | fixes x y::"'a::at_base" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2209 | shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2210 | and "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2211 | and "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2212 | unfolding atom_eq_iff [symmetric] | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2213 | unfolding atom_eqvt [symmetric] | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2214 | by simp_all | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2215 | |
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2216 | lemma obtain_at_base: | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2217 | assumes X: "finite X" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2218 | obtains a::"'a::at_base" where "atom a \<notin> X" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2219 | proof - | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2220 | have "inj (atom :: 'a \<Rightarrow> atom)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2221 | by (simp add: inj_on_def) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2222 | with X have "finite (atom -` X :: 'a set)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2223 | by (rule finite_vimageI) | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2224 | with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2225 | by auto | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2226 | then obtain a :: 'a where "atom a \<notin> X" | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2227 | by auto | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2228 | thus ?thesis .. | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2229 | qed | 
| 
84a13d1e2511
moved mk_atom into the library; that meant that concrete atom classes need to be in Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1941diff
changeset | 2230 | |
| 1973 
fc5ce7f22b74
use the more general type-class at_base
 Christian Urban <urbanc@in.tum.de> parents: 
1972diff
changeset | 2231 | lemma supp_finite_set_at_base: | 
| 1971 
8daf6ff5e11a
simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1962diff
changeset | 2232 | assumes a: "finite S" | 
| 
8daf6ff5e11a
simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1962diff
changeset | 2233 | shows "supp S = atom ` S" | 
| 2565 
6bf332360510
moved most material fron Nominal2_FSet into the Nominal_Base theory
 Christian Urban <urbanc@in.tum.de> parents: 
2560diff
changeset | 2234 | apply(simp add: supp_of_finite_sets[OF a]) | 
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 2235 | apply(simp add: supp_at_base) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 2236 | apply(auto) | 
| 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 2237 | done | 
| 1971 
8daf6ff5e11a
simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1962diff
changeset | 2238 | |
| 2657 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2239 | lemma fresh_finite_set_at_base: | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2240 | fixes a::"'a::at_base" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2241 | assumes a: "finite S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2242 | shows "atom a \<sharp> S \<longleftrightarrow> a \<notin> S" | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2243 | unfolding fresh_def | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2244 | apply(simp add: supp_finite_set_at_base[OF a]) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2245 | apply(subst inj_image_mem_iff) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2246 | apply(simp add: inj_on_def) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2247 | apply(simp) | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2248 | done | 
| 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2249 | |
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2250 | lemma fresh_at_base_permute_iff[simp]: | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2251 | fixes a::"'a::at_base" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2252 | shows "atom (p \<bullet> a) \<sharp> p \<bullet> x \<longleftrightarrow> atom a \<sharp> x" | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2253 | unfolding atom_eqvt[symmetric] | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2254 | by (simp add: fresh_permute_iff) | 
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2255 | |
| 2657 
1ea9c059fc0f
a few lemmas about freshness for at and at_base
 Christian Urban <urbanc@in.tum.de> parents: 
2641diff
changeset | 2256 | |
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2257 | section {* Infrastructure for concrete atom types *}
 | 
| 1971 
8daf6ff5e11a
simpliied and moved the remaining lemmas about the atom-function to Nominal2_Base
 Christian Urban <urbanc@in.tum.de> parents: 
1962diff
changeset | 2258 | |
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2259 | definition | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2260 |   flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2261 | where | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2262 | "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2263 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2264 | lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2265 | unfolding flip_def by (rule swap_self) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2266 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2267 | lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2268 | unfolding flip_def by (rule swap_commute) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2269 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2270 | lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2271 | unfolding flip_def by (rule minus_swap) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2272 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2273 | lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2274 | unfolding flip_def by (rule swap_cancel) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2275 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2276 | lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2277 | unfolding permute_plus [symmetric] add_flip_cancel by simp | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2278 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2279 | lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2280 | by (simp add: flip_commute) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2281 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2282 | lemma flip_eqvt: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2283 | fixes a b c::"'a::at_base" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2284 | shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2285 | unfolding flip_def | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2286 | by (simp add: swap_eqvt atom_eqvt) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2287 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2288 | lemma flip_at_base_simps [simp]: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2289 | shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2290 | and "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2291 | and "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2292 | and "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2293 | unfolding flip_def | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2294 | unfolding atom_eq_iff [symmetric] | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2295 | unfolding atom_eqvt [symmetric] | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2296 | by simp_all | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2297 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2298 | text {* the following two lemmas do not hold for at_base, 
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2299 | only for single sort atoms from at *} | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2300 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2301 | lemma permute_flip_at: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2302 | fixes a b c::"'a::at" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2303 | shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2304 | unfolding flip_def | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2305 | apply (rule atom_eq_iff [THEN iffD1]) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2306 | apply (subst atom_eqvt [symmetric]) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2307 | apply (simp add: swap_atom) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2308 | done | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2309 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2310 | lemma flip_at_simps [simp]: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2311 | fixes a b::"'a::at" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2312 | shows "(a \<leftrightarrow> b) \<bullet> a = b" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2313 | and "(a \<leftrightarrow> b) \<bullet> b = a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2314 | unfolding permute_flip_at by simp_all | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2315 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2316 | lemma flip_fresh_fresh: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2317 | fixes a b::"'a::at_base" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2318 | assumes "atom a \<sharp> x" "atom b \<sharp> x" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2319 | shows "(a \<leftrightarrow> b) \<bullet> x = x" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2320 | using assms | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2321 | by (simp add: flip_def swap_fresh_fresh) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2322 | |
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2323 | |
| 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2324 | |
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2325 | subsection {* Syntax for coercing at-elements to the atom-type *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2326 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2327 | syntax | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2328 |   "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2329 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2330 | translations | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2331 | "_atom_constrain a t" => "CONST atom (_constrain a t)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2332 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2333 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2334 | subsection {* A lemma for proving instances of class @{text at}. *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2335 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2336 | setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2337 | setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2338 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2339 | text {*
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2340 |   New atom types are defined as subtypes of @{typ atom}.
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2341 | *} | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2342 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2343 | lemma exists_eq_simple_sort: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2344 |   shows "\<exists>a. a \<in> {a. sort_of a = s}"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2345 | by (rule_tac x="Atom s 0" in exI, simp) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2346 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2347 | lemma exists_eq_sort: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2348 |   shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2349 | by (rule_tac x="Atom (sort_fun x) y" in exI, simp) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2350 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2351 | lemma at_base_class: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2352 | fixes sort_fun :: "'b \<Rightarrow>atom_sort" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2353 | fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2354 |   assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2355 | assumes atom_def: "\<And>a. atom a = Rep a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2356 | assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2357 |   shows "OFCLASS('a, at_base_class)"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2358 | proof | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2359 |   interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2360 | have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2361 | fix a b :: 'a and p p1 p2 :: perm | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2362 | show "0 \<bullet> a = a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2363 | unfolding permute_def by (simp add: Rep_inverse) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2364 | show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2365 | unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2366 | show "atom a = atom b \<longleftrightarrow> a = b" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2367 | unfolding atom_def by (simp add: Rep_inject) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2368 | show "p \<bullet> atom a = atom (p \<bullet> a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2369 | unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2370 | qed | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2371 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2372 | (* | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2373 | lemma at_class: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2374 | fixes s :: atom_sort | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2375 | fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2376 |   assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2377 | assumes atom_def: "\<And>a. atom a = Rep a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2378 | assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2379 |   shows "OFCLASS('a, at_class)"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2380 | proof | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2381 |   interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2382 | have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2383 | fix a b :: 'a and p p1 p2 :: perm | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2384 | show "0 \<bullet> a = a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2385 | unfolding permute_def by (simp add: Rep_inverse) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2386 | show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2387 | unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2388 | show "sort_of (atom a) = sort_of (atom b)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2389 | unfolding atom_def by (simp add: sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2390 | show "atom a = atom b \<longleftrightarrow> a = b" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2391 | unfolding atom_def by (simp add: Rep_inject) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2392 | show "p \<bullet> atom a = atom (p \<bullet> a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2393 | unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2394 | qed | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2395 | *) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2396 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2397 | lemma at_class: | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2398 | fixes s :: atom_sort | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2399 | fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2400 |   assumes type: "type_definition Rep Abs {a. sort_of a = s}"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2401 | assumes atom_def: "\<And>a. atom a = Rep a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2402 | assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2403 |   shows "OFCLASS('a, at_class)"
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2404 | proof | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2405 |   interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2406 | have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2407 | fix a b :: 'a and p p1 p2 :: perm | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2408 | show "0 \<bullet> a = a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2409 | unfolding permute_def by (simp add: Rep_inverse) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2410 | show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2411 | unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2412 | show "sort_of (atom a) = sort_of (atom b)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2413 | unfolding atom_def by (simp add: sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2414 | show "atom a = atom b \<longleftrightarrow> a = b" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2415 | unfolding atom_def by (simp add: Rep_inject) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2416 | show "p \<bullet> atom a = atom (p \<bullet> a)" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2417 | unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2418 | qed | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2419 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2420 | setup {* Sign.add_const_constraint
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2421 |   (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2422 | setup {* Sign.add_const_constraint
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2423 |   (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2424 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2425 | |
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2426 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2427 | section {* The freshness lemma according to Andy Pitts *}
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2428 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2429 | lemma freshness_lemma: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2430 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2431 | assumes a: "\<exists>a. atom a \<sharp> (h, h a)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2432 | shows "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2433 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2434 | from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2435 | by (auto simp add: fresh_Pair) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2436 | show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2437 | proof (intro exI allI impI) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2438 | fix a :: 'a | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2439 | assume a3: "atom a \<sharp> h" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2440 | show "h a = h b" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2441 | proof (cases "a = b") | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2442 | assume "a = b" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2443 | thus "h a = h b" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2444 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2445 | assume "a \<noteq> b" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2446 | hence "atom a \<sharp> b" by (simp add: fresh_at_base) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2447 | with a3 have "atom a \<sharp> h b" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2448 | by (rule fresh_fun_app) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2449 | with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2450 | by (rule swap_fresh_fresh) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2451 | from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2452 | by (rule swap_fresh_fresh) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2453 | from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2454 | also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2455 | by (rule permute_fun_app_eq) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2456 | also have "\<dots> = h a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2457 | using d2 by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2458 | finally show "h a = h b" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2459 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2460 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2461 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2462 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2463 | lemma freshness_lemma_unique: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2464 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2465 | assumes a: "\<exists>a. atom a \<sharp> (h, h a)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2466 | shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2467 | proof (rule ex_ex1I) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2468 | from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2469 | by (rule freshness_lemma) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2470 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2471 | fix x y | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2472 | assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2473 | assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2474 | from a x y show "x = y" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2475 | by (auto simp add: fresh_Pair) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2476 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2477 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2478 | text {* packaging the freshness lemma into a function *}
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2479 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2480 | definition | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2481 |   fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
 | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2482 | where | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2483 | "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2484 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2485 | lemma fresh_fun_apply: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2486 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2487 | assumes a: "\<exists>a. atom a \<sharp> (h, h a)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2488 | assumes b: "atom a \<sharp> h" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2489 | shows "fresh_fun h = h a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2490 | unfolding fresh_fun_def | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2491 | proof (rule the_equality) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2492 | show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2493 | proof (intro strip) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2494 | fix a':: 'a | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2495 | assume c: "atom a' \<sharp> h" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2496 | from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2497 | with b c show "h a' = h a" by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2498 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2499 | next | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2500 | fix fr :: 'b | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2501 | assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2502 | with b show "fr = h a" by auto | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2503 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2504 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2505 | lemma fresh_fun_apply': | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2506 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2507 | assumes a: "atom a \<sharp> h" "atom a \<sharp> h a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2508 | shows "fresh_fun h = h a" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2509 | apply (rule fresh_fun_apply) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2510 | apply (auto simp add: fresh_Pair intro: a) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2511 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2512 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2513 | lemma fresh_fun_eqvt: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2514 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2515 | assumes a: "\<exists>a. atom a \<sharp> (h, h a)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2516 | shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2517 | using a | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2518 | apply (clarsimp simp add: fresh_Pair) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2519 | apply (subst fresh_fun_apply', assumption+) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2520 | apply (drule fresh_permute_iff [where p=p, THEN iffD2]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2521 | apply (drule fresh_permute_iff [where p=p, THEN iffD2]) | 
| 2683 
42c0d011a177
ported some of the old proofs to serve as testcases
 Christian Urban <urbanc@in.tum.de> parents: 
2679diff
changeset | 2522 | apply (simp only: atom_eqvt permute_fun_app_eq [where f=h]) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2523 | apply (erule (1) fresh_fun_apply' [symmetric]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2524 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2525 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2526 | lemma fresh_fun_supports: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2527 | fixes h :: "'a::at \<Rightarrow> 'b::pt" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2528 | assumes a: "\<exists>a. atom a \<sharp> (h, h a)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2529 | shows "(supp h) supports (fresh_fun h)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2530 | apply (simp add: supports_def fresh_def [symmetric]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2531 | apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2532 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2533 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2534 | notation fresh_fun (binder "FRESH " 10) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2535 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2536 | lemma FRESH_f_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2537 | fixes P :: "'a::at \<Rightarrow> 'b::pure" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2538 | fixes f :: "'b \<Rightarrow> 'c::pure" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2539 | assumes P: "finite (supp P)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2540 | shows "(FRESH x. f (P x)) = f (FRESH x. P x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2541 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2542 | obtain a::'a where "atom a \<notin> supp P" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2543 | using P by (rule obtain_at_base) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2544 | hence "atom a \<sharp> P" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2545 | by (simp add: fresh_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2546 | show "(FRESH x. f (P x)) = f (FRESH x. P x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2547 | apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2548 | apply (cut_tac `atom a \<sharp> P`) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2549 | apply (simp add: fresh_conv_MOST) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2550 | apply (elim MOST_rev_mp, rule MOST_I, clarify) | 
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 2551 | apply (simp add: permute_fun_def permute_pure fun_eq_iff) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2552 | apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2553 | apply (rule refl) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2554 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2555 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2556 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2557 | lemma FRESH_binop_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2558 | fixes P :: "'a::at \<Rightarrow> 'b::pure" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2559 | fixes Q :: "'a::at \<Rightarrow> 'c::pure" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2560 | fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2561 | assumes P: "finite (supp P)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2562 | and Q: "finite (supp Q)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2563 | shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2564 | proof - | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2565 | from assms have "finite (supp P \<union> supp Q)" by simp | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2566 | then obtain a::'a where "atom a \<notin> (supp P \<union> supp Q)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2567 | by (rule obtain_at_base) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2568 | hence "atom a \<sharp> P" and "atom a \<sharp> Q" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2569 | by (simp_all add: fresh_def) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2570 | show ?thesis | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2571 | apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2572 | apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2573 | apply (simp add: fresh_conv_MOST) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2574 | apply (elim MOST_rev_mp, rule MOST_I, clarify) | 
| 2479 
a9b6a00b1ba0
updated to Isabelle Sept 16
 Christian Urban <urbanc@in.tum.de> parents: 
2475diff
changeset | 2575 | apply (simp add: permute_fun_def permute_pure fun_eq_iff) | 
| 2470 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2576 | apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2577 | apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh]) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2578 | apply (rule refl) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2579 | done | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2580 | qed | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2581 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2582 | lemma FRESH_conj_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2583 | fixes P Q :: "'a::at \<Rightarrow> bool" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2584 | assumes P: "finite (supp P)" and Q: "finite (supp Q)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2585 | shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2586 | using P Q by (rule FRESH_binop_iff) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2587 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2588 | lemma FRESH_disj_iff: | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2589 | fixes P Q :: "'a::at \<Rightarrow> bool" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2590 | assumes P: "finite (supp P)" and Q: "finite (supp Q)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2591 | shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)" | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2592 | using P Q by (rule FRESH_binop_iff) | 
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2593 | |
| 
bdb1eab47161
moved everything out of Nominal_Supp
 Christian Urban <urbanc@in.tum.de> parents: 
2467diff
changeset | 2594 | |
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2595 | section {* Library functions for the nominal infrastructure *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2596 | |
| 1833 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1774diff
changeset | 2597 | use "nominal_library.ML" | 
| 
2050b5723c04
added a library for basic nominal functions; separated nominal_eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1774diff
changeset | 2598 | |
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 2599 | |
| 2467 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2600 | section {* Automation for creating concrete atom types *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2601 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2602 | text {* at the moment only single-sort concrete atoms are supported *}
 | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2603 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2604 | use "nominal_atoms.ML" | 
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2605 | |
| 
67b3933c3190
got rid of Nominal_Atoms (folded into Nominal2_Base)
 Christian Urban <urbanc@in.tum.de> parents: 
2466diff
changeset | 2606 | |
| 2466 
47c840599a6b
cleaned a bit various thy-files in Nominal-General
 Christian Urban <urbanc@in.tum.de> parents: 
2378diff
changeset | 2607 | |
| 1062 | 2608 | end |