|
1 (* Title: Nominal2_Base |
|
2 Authors: Christian Urban, Brian Huffman, Cezary Kaliszyk |
|
3 |
|
4 Basic definitions and lemma infrastructure for |
|
5 Nominal Isabelle. |
|
6 *) |
|
7 theory Nominal2_Base |
|
8 imports Main |
|
9 "~~/src/HOL/Library/Infinite_Set" |
|
10 "~~/src/HOL/Quotient_Examples/FSet" |
|
11 "GPerm" |
|
12 keywords |
|
13 "atom_decl" "equivariance" :: thy_decl |
|
14 uses ("nominal_basics.ML") |
|
15 ("nominal_thmdecls.ML") |
|
16 ("nominal_permeq.ML") |
|
17 ("nominal_library.ML") |
|
18 ("nominal_atoms.ML") |
|
19 ("nominal_eqvt.ML") |
|
20 begin |
|
21 |
|
22 section {* Atoms and Sorts *} |
|
23 |
|
24 text {* A simple implementation for atom_sorts is strings. *} |
|
25 (* types atom_sort = string *) |
|
26 |
|
27 text {* To deal with Church-like binding we use trees of |
|
28 strings as sorts. *} |
|
29 |
|
30 datatype atom_sort = Sort "string" "atom_sort list" |
|
31 |
|
32 datatype atom = Atom atom_sort nat |
|
33 |
|
34 |
|
35 text {* Basic projection function. *} |
|
36 |
|
37 primrec |
|
38 sort_of :: "atom \<Rightarrow> atom_sort" |
|
39 where |
|
40 "sort_of (Atom s n) = s" |
|
41 |
|
42 primrec |
|
43 nat_of :: "atom \<Rightarrow> nat" |
|
44 where |
|
45 "nat_of (Atom s n) = n" |
|
46 |
|
47 |
|
48 text {* There are infinitely many atoms of each sort. *} |
|
49 lemma INFM_sort_of_eq: |
|
50 shows "INFM a. sort_of a = s" |
|
51 proof - |
|
52 have "INFM i. sort_of (Atom s i) = s" by simp |
|
53 moreover have "inj (Atom s)" by (simp add: inj_on_def) |
|
54 ultimately show "INFM a. sort_of a = s" by (rule INFM_inj) |
|
55 qed |
|
56 |
|
57 lemma infinite_sort_of_eq: |
|
58 shows "infinite {a. sort_of a = s}" |
|
59 using INFM_sort_of_eq unfolding INFM_iff_infinite . |
|
60 |
|
61 lemma atom_infinite [simp]: |
|
62 shows "infinite (UNIV :: atom set)" |
|
63 using subset_UNIV infinite_sort_of_eq |
|
64 by (rule infinite_super) |
|
65 |
|
66 lemma obtain_atom: |
|
67 fixes X :: "atom set" |
|
68 assumes X: "finite X" |
|
69 obtains a where "a \<notin> X" "sort_of a = s" |
|
70 proof - |
|
71 from X have "MOST a. a \<notin> X" |
|
72 unfolding MOST_iff_cofinite by simp |
|
73 with INFM_sort_of_eq |
|
74 have "INFM a. sort_of a = s \<and> a \<notin> X" |
|
75 by (rule INFM_conjI) |
|
76 then obtain a where "a \<notin> X" "sort_of a = s" |
|
77 by (auto elim: INFM_E) |
|
78 then show ?thesis .. |
|
79 qed |
|
80 |
|
81 lemma atom_components_eq_iff: |
|
82 fixes a b :: atom |
|
83 shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b" |
|
84 by (induct a, induct b, simp) |
|
85 |
|
86 |
|
87 section {* Sort-Respecting Permutations *} |
|
88 |
|
89 definition "sort_respecting p \<longleftrightarrow> (\<forall>a. sort_of (gpermute p a) = sort_of a)" |
|
90 |
|
91 lemma sort_respecting_0[simp]: |
|
92 "sort_respecting (0\<Colon>atom gperm)" |
|
93 by (simp add: sort_respecting_def) |
|
94 |
|
95 typedef (open) perm = "{p::atom gperm. sort_respecting p}" |
|
96 by (auto intro: exI[of _ "0"]) |
|
97 |
|
98 lemma perm_eq_rep: |
|
99 "p = q \<longleftrightarrow> Rep_perm p = Rep_perm q" |
|
100 by (simp add: Rep_perm_inject) |
|
101 |
|
102 definition mk_perm :: "atom gperm \<Rightarrow> perm" where |
|
103 "mk_perm p = Abs_perm (if sort_respecting p then p else 0)" |
|
104 |
|
105 lemma sort_respecting_Rep_perm [simp, intro]: |
|
106 "sort_respecting (Rep_perm p)" |
|
107 using Rep_perm [of p] by simp |
|
108 |
|
109 lemma Rep_perm_mk_perm [simp]: |
|
110 "Rep_perm (mk_perm p) = (if sort_respecting p then p else 0)" |
|
111 by (simp add: mk_perm_def Abs_perm_inverse) |
|
112 |
|
113 lemma mk_perm_Rep_perm [simp, code abstype]: |
|
114 "mk_perm (Rep_perm dxs) = dxs" |
|
115 by (simp add: mk_perm_def Rep_perm_inverse) |
|
116 |
|
117 instance perm :: size .. |
|
118 |
|
119 instantiation perm :: group_add |
|
120 begin |
|
121 |
|
122 definition "(0 :: perm) = mk_perm 0" |
|
123 |
|
124 definition "uminus p = mk_perm (uminus (Rep_perm p))" |
|
125 |
|
126 definition "p + q = mk_perm ((Rep_perm p) + (Rep_perm q))" |
|
127 |
|
128 definition "(p :: perm) - q = p + - q" |
|
129 |
|
130 lemma [simp]: |
|
131 "sort_respecting x \<Longrightarrow> sort_respecting y \<Longrightarrow> sort_respecting (x + y)" |
|
132 unfolding sort_respecting_def |
|
133 by descending (simp add: perm_add_apply) |
|
134 |
|
135 lemma [simp]: |
|
136 "sort_respecting y \<Longrightarrow> sort_respecting (- y)" |
|
137 unfolding sort_respecting_def |
|
138 by partiality_descending |
|
139 (auto, metis perm_apply_minus) |
|
140 |
|
141 lemma Rep_perm_0 [simp, code abstract]: |
|
142 "Rep_perm 0 = 0" |
|
143 by (simp add: zero_perm_def) |
|
144 |
|
145 lemma Rep_perm_uminus [simp, code abstract]: |
|
146 "Rep_perm (- p) = - (Rep_perm p)" |
|
147 by (simp add: uminus_perm_def) |
|
148 |
|
149 lemma Rep_perm_add [simp, code abstract]: |
|
150 "Rep_perm (p + q) = (Rep_perm p) + (Rep_perm q)" |
|
151 by (simp add: plus_perm_def) |
|
152 |
|
153 instance |
|
154 by default (auto simp add: perm_eq_rep add_assoc minus_perm_def) |
|
155 |
|
156 end |
|
157 |
|
158 definition swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')") |
|
159 where "swap a b = (if sort_of a = sort_of b then mk_perm (gswap a b) else 0)" |
|
160 |
|
161 lemma sort_respecting_swap [simp]: |
|
162 "sort_of a = sort_of b \<Longrightarrow> sort_respecting (gswap a b)" |
|
163 unfolding sort_respecting_def |
|
164 by descending auto |
|
165 |
|
166 lemma Rep_swap [simp, code abstract]: |
|
167 "Rep_perm (swap a b) = (if sort_of a = sort_of b then gswap a b else 0)" |
|
168 by (simp add: swap_def) |
|
169 |
|
170 lemma swap_different_sorts [simp]: |
|
171 "sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0" |
|
172 by (simp add: perm_eq_rep) |
|
173 |
|
174 lemma swap_cancel: |
|
175 shows "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0" |
|
176 and "(a \<rightleftharpoons> b) + (b \<rightleftharpoons> a) = 0" |
|
177 by (simp_all add: perm_eq_rep) |
|
178 |
|
179 lemma swap_self [simp]: |
|
180 "(a \<rightleftharpoons> a) = 0" |
|
181 by (simp add: perm_eq_rep) |
|
182 |
|
183 lemma minus_swap [simp]: |
|
184 "- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)" |
|
185 by (simp add: perm_eq_rep) |
|
186 |
|
187 lemma swap_commute: |
|
188 "(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)" |
|
189 by (simp add: perm_eq_rep swap_commute) |
|
190 |
|
191 lemma swap_triple: |
|
192 assumes "a \<noteq> b" and "c \<noteq> b" |
|
193 assumes "sort_of a = sort_of b" "sort_of b = sort_of c" |
|
194 shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)" |
|
195 using assms by (simp add: perm_eq_rep swap_triple) |
|
196 |
|
197 section {* Permutation Types *} |
|
198 |
|
199 text {* |
|
200 Infix syntax for @{text permute} has higher precedence than |
|
201 addition, but lower than unary minus. |
|
202 *} |
|
203 |
|
204 class pt = |
|
205 fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75) |
|
206 assumes permute_zero [simp]: "0 \<bullet> x = x" |
|
207 assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)" |
|
208 begin |
|
209 |
|
210 lemma permute_diff [simp]: |
|
211 shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x" |
|
212 unfolding diff_minus by simp |
|
213 |
|
214 lemma permute_minus_cancel [simp]: |
|
215 shows "p \<bullet> - p \<bullet> x = x" |
|
216 and "- p \<bullet> p \<bullet> x = x" |
|
217 unfolding permute_plus [symmetric] by simp_all |
|
218 |
|
219 lemma permute_swap_cancel [simp]: |
|
220 shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x" |
|
221 unfolding permute_plus [symmetric] |
|
222 by (simp add: swap_cancel) |
|
223 |
|
224 lemma permute_swap_cancel2 [simp]: |
|
225 shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x" |
|
226 unfolding permute_plus [symmetric] |
|
227 by (simp add: swap_commute) |
|
228 |
|
229 lemma inj_permute [simp]: |
|
230 shows "inj (permute p)" |
|
231 by (rule inj_on_inverseI) |
|
232 (rule permute_minus_cancel) |
|
233 |
|
234 lemma surj_permute [simp]: |
|
235 shows "surj (permute p)" |
|
236 by (rule surjI, rule permute_minus_cancel) |
|
237 |
|
238 lemma bij_permute [simp]: |
|
239 shows "bij (permute p)" |
|
240 by (rule bijI [OF inj_permute surj_permute]) |
|
241 |
|
242 lemma inv_permute: |
|
243 shows "inv (permute p) = permute (- p)" |
|
244 by (rule inv_equality) (simp_all) |
|
245 |
|
246 lemma permute_minus: |
|
247 shows "permute (- p) = inv (permute p)" |
|
248 by (simp add: inv_permute) |
|
249 |
|
250 lemma permute_eq_iff [simp]: |
|
251 shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y" |
|
252 by (rule inj_permute [THEN inj_eq]) |
|
253 |
|
254 end |
|
255 |
|
256 subsection {* Permutations for atoms *} |
|
257 |
|
258 instantiation atom :: pt |
|
259 begin |
|
260 |
|
261 definition |
|
262 "p \<bullet> a = gpermute (Rep_perm p) a" |
|
263 |
|
264 instance |
|
265 by default (simp_all add: permute_atom_def) |
|
266 |
|
267 end |
|
268 |
|
269 lemma sort_of_permute [simp]: |
|
270 shows "sort_of (p \<bullet> a) = sort_of a" |
|
271 by (metis sort_respecting_Rep_perm sort_respecting_def permute_atom_def) |
|
272 |
|
273 lemma swap_atom: |
|
274 shows "(a \<rightleftharpoons> b) \<bullet> c = |
|
275 (if sort_of a = sort_of b |
|
276 then (if c = a then b else if c = b then a else c) else c)" |
|
277 by (auto simp add: permute_atom_def) |
|
278 |
|
279 lemma swap_atom_simps [simp]: |
|
280 "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b" |
|
281 "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a" |
|
282 "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c" |
|
283 unfolding swap_atom by simp_all |
|
284 |
|
285 lemma perm_eq_iff: |
|
286 fixes p q :: "perm" |
|
287 shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)" |
|
288 unfolding permute_atom_def perm_eq_rep |
|
289 by (simp add: gperm_eq) |
|
290 |
|
291 subsection {* Permutations for permutations *} |
|
292 |
|
293 instantiation perm :: pt |
|
294 begin |
|
295 |
|
296 definition |
|
297 "p \<bullet> q = p + q - p" |
|
298 |
|
299 instance |
|
300 by default |
|
301 (simp_all add: permute_perm_def diff_minus minus_add add_assoc) |
|
302 |
|
303 end |
|
304 |
|
305 lemma permute_self: |
|
306 shows "p \<bullet> p = p" |
|
307 unfolding permute_perm_def |
|
308 by (simp add: diff_minus add_assoc) |
|
309 |
|
310 lemma pemute_minus_self: |
|
311 shows "- p \<bullet> p = p" |
|
312 unfolding permute_perm_def |
|
313 by (simp add: diff_minus add_assoc) |
|
314 |
|
315 |
|
316 subsection {* Permutations for functions *} |
|
317 |
|
318 instantiation "fun" :: (pt, pt) pt |
|
319 begin |
|
320 |
|
321 definition |
|
322 "p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))" |
|
323 |
|
324 instance |
|
325 by default |
|
326 (simp_all add: permute_fun_def minus_add) |
|
327 |
|
328 end |
|
329 |
|
330 lemma permute_fun_app_eq: |
|
331 shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)" |
|
332 unfolding permute_fun_def by simp |
|
333 |
|
334 |
|
335 subsection {* Permutations for booleans *} |
|
336 |
|
337 instantiation bool :: pt |
|
338 begin |
|
339 |
|
340 definition "p \<bullet> (b::bool) = b" |
|
341 |
|
342 instance |
|
343 by (default) |
|
344 (simp_all add: permute_bool_def) |
|
345 |
|
346 end |
|
347 |
|
348 lemma permute_boolE: |
|
349 fixes P::"bool" |
|
350 shows "p \<bullet> P \<Longrightarrow> P" |
|
351 by (simp add: permute_bool_def) |
|
352 |
|
353 lemma permute_boolI: |
|
354 fixes P::"bool" |
|
355 shows "P \<Longrightarrow> p \<bullet> P" |
|
356 by(simp add: permute_bool_def) |
|
357 |
|
358 subsection {* Permutations for sets *} |
|
359 |
|
360 instantiation "set" :: (pt) pt |
|
361 begin |
|
362 |
|
363 definition |
|
364 "p \<bullet> X = {p \<bullet> x | x. x \<in> X}" |
|
365 |
|
366 instance |
|
367 apply default |
|
368 apply (auto simp add: permute_set_def) |
|
369 done |
|
370 |
|
371 end |
|
372 |
|
373 lemma permute_set_eq: |
|
374 shows "p \<bullet> X = {x. - p \<bullet> x \<in> X}" |
|
375 unfolding permute_set_def |
|
376 by (auto) (metis permute_minus_cancel(1)) |
|
377 |
|
378 lemma permute_set_eq_image: |
|
379 shows "p \<bullet> X = permute p ` X" |
|
380 unfolding permute_set_def by auto |
|
381 |
|
382 lemma permute_set_eq_vimage: |
|
383 shows "p \<bullet> X = permute (- p) -` X" |
|
384 unfolding permute_set_eq vimage_def |
|
385 by simp |
|
386 |
|
387 lemma permute_finite [simp]: |
|
388 shows "finite (p \<bullet> X) = finite X" |
|
389 unfolding permute_set_eq_vimage |
|
390 using bij_permute by (rule finite_vimage_iff) |
|
391 |
|
392 lemma swap_set_not_in: |
|
393 assumes a: "a \<notin> S" "b \<notin> S" |
|
394 shows "(a \<rightleftharpoons> b) \<bullet> S = S" |
|
395 unfolding permute_set_def |
|
396 using a by (auto simp add: swap_atom) |
|
397 |
|
398 lemma swap_set_in: |
|
399 assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b" |
|
400 shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S" |
|
401 unfolding permute_set_def |
|
402 using a by (auto simp add: swap_atom) |
|
403 |
|
404 lemma swap_set_in_eq: |
|
405 assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b" |
|
406 shows "(a \<rightleftharpoons> b) \<bullet> S = (S - {a}) \<union> {b}" |
|
407 unfolding permute_set_def |
|
408 using a by (auto simp add: swap_atom) |
|
409 |
|
410 lemma swap_set_both_in: |
|
411 assumes a: "a \<in> S" "b \<in> S" |
|
412 shows "(a \<rightleftharpoons> b) \<bullet> S = S" |
|
413 unfolding permute_set_def |
|
414 using a by (auto simp add: swap_atom) |
|
415 |
|
416 lemma mem_permute_iff: |
|
417 shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X" |
|
418 unfolding permute_set_def |
|
419 by auto |
|
420 |
|
421 lemma empty_eqvt: |
|
422 shows "p \<bullet> {} = {}" |
|
423 unfolding permute_set_def |
|
424 by (simp) |
|
425 |
|
426 lemma insert_eqvt: |
|
427 shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)" |
|
428 unfolding permute_set_eq_image image_insert .. |
|
429 |
|
430 |
|
431 subsection {* Permutations for @{typ unit} *} |
|
432 |
|
433 instantiation unit :: pt |
|
434 begin |
|
435 |
|
436 definition "p \<bullet> (u::unit) = u" |
|
437 |
|
438 instance |
|
439 by (default) (simp_all add: permute_unit_def) |
|
440 |
|
441 end |
|
442 |
|
443 |
|
444 subsection {* Permutations for products *} |
|
445 |
|
446 instantiation prod :: (pt, pt) pt |
|
447 begin |
|
448 |
|
449 primrec |
|
450 permute_prod |
|
451 where |
|
452 Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)" |
|
453 |
|
454 instance |
|
455 by default auto |
|
456 |
|
457 end |
|
458 |
|
459 subsection {* Permutations for sums *} |
|
460 |
|
461 instantiation sum :: (pt, pt) pt |
|
462 begin |
|
463 |
|
464 primrec |
|
465 permute_sum |
|
466 where |
|
467 Inl_eqvt: "p \<bullet> (Inl x) = Inl (p \<bullet> x)" |
|
468 | Inr_eqvt: "p \<bullet> (Inr y) = Inr (p \<bullet> y)" |
|
469 |
|
470 instance |
|
471 by (default) (case_tac [!] x, simp_all) |
|
472 |
|
473 end |
|
474 |
|
475 subsection {* Permutations for @{typ "'a list"} *} |
|
476 |
|
477 instantiation list :: (pt) pt |
|
478 begin |
|
479 |
|
480 primrec |
|
481 permute_list |
|
482 where |
|
483 Nil_eqvt: "p \<bullet> [] = []" |
|
484 | Cons_eqvt: "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs" |
|
485 |
|
486 instance |
|
487 by (default) (induct_tac [!] x, simp_all) |
|
488 |
|
489 end |
|
490 |
|
491 lemma set_eqvt: |
|
492 shows "p \<bullet> (set xs) = set (p \<bullet> xs)" |
|
493 by (induct xs) (simp_all add: empty_eqvt insert_eqvt) |
|
494 |
|
495 |
|
496 |
|
497 subsection {* Permutations for @{typ "'a option"} *} |
|
498 |
|
499 instantiation option :: (pt) pt |
|
500 begin |
|
501 |
|
502 primrec |
|
503 permute_option |
|
504 where |
|
505 None_eqvt: "p \<bullet> None = None" |
|
506 | Some_eqvt: "p \<bullet> (Some x) = Some (p \<bullet> x)" |
|
507 |
|
508 instance |
|
509 by (default) (induct_tac [!] x, simp_all) |
|
510 |
|
511 end |
|
512 |
|
513 subsection {* Permutations for @{typ "'a multiset"} *} |
|
514 |
|
515 instantiation multiset :: (pt) pt |
|
516 begin |
|
517 |
|
518 definition |
|
519 "p \<bullet> M = {# p \<bullet> x. x :# M #}" |
|
520 |
|
521 instance |
|
522 proof |
|
523 fix M :: "'a multiset" and p q :: "perm" |
|
524 show "0 \<bullet> M = M" |
|
525 unfolding permute_multiset_def |
|
526 by (induct_tac M) (simp_all) |
|
527 show "(p + q) \<bullet> M = p \<bullet> q \<bullet> M" |
|
528 unfolding permute_multiset_def |
|
529 by (induct_tac M) (simp_all) |
|
530 qed |
|
531 |
|
532 end |
|
533 |
|
534 lemma permute_multiset [simp]: |
|
535 fixes M N::"('a::pt) multiset" |
|
536 shows "(p \<bullet> {#}) = ({#} ::('a::pt) multiset)" |
|
537 and "(p \<bullet> {# x #}) = {# p \<bullet> x #}" |
|
538 and "(p \<bullet> (M + N)) = (p \<bullet> M) + (p \<bullet> N)" |
|
539 unfolding permute_multiset_def |
|
540 by (simp_all) |
|
541 |
|
542 |
|
543 subsection {* Permutations for @{typ "'a fset"} *} |
|
544 |
|
545 lemma permute_fset_rsp[quot_respect]: |
|
546 shows "(op = ===> list_eq ===> list_eq) permute permute" |
|
547 unfolding fun_rel_def |
|
548 by (simp add: set_eqvt[symmetric]) |
|
549 |
|
550 instantiation fset :: (pt) pt |
|
551 begin |
|
552 |
|
553 quotient_definition |
|
554 "permute_fset :: perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
555 is |
|
556 "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
557 |
|
558 instance |
|
559 proof |
|
560 fix x :: "'a fset" and p q :: "perm" |
|
561 have lst: "\<And>l :: 'a list. 0 \<bullet> l = l" by simp |
|
562 show "0 \<bullet> x = x" by (lifting lst) |
|
563 have lst: "\<And>p q :: perm. \<And>x :: 'a list. (p + q) \<bullet> x = p \<bullet> q \<bullet> x" by simp |
|
564 show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" by (lifting lst) |
|
565 qed |
|
566 |
|
567 end |
|
568 |
|
569 lemma permute_fset [simp]: |
|
570 fixes S::"('a::pt) fset" |
|
571 shows "(p \<bullet> {||}) = ({||} ::('a::pt) fset)" |
|
572 and "(p \<bullet> insert_fset x S) = insert_fset (p \<bullet> x) (p \<bullet> S)" |
|
573 by (lifting permute_list.simps) |
|
574 |
|
575 lemma fset_eqvt: |
|
576 shows "p \<bullet> (fset S) = fset (p \<bullet> S)" |
|
577 by (lifting set_eqvt) |
|
578 |
|
579 |
|
580 subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *} |
|
581 |
|
582 instantiation char :: pt |
|
583 begin |
|
584 |
|
585 definition "p \<bullet> (c::char) = c" |
|
586 |
|
587 instance |
|
588 by (default) (simp_all add: permute_char_def) |
|
589 |
|
590 end |
|
591 |
|
592 instantiation nat :: pt |
|
593 begin |
|
594 |
|
595 definition "p \<bullet> (n::nat) = n" |
|
596 |
|
597 instance |
|
598 by (default) (simp_all add: permute_nat_def) |
|
599 |
|
600 end |
|
601 |
|
602 instantiation int :: pt |
|
603 begin |
|
604 |
|
605 definition "p \<bullet> (i::int) = i" |
|
606 |
|
607 instance |
|
608 by (default) (simp_all add: permute_int_def) |
|
609 |
|
610 end |
|
611 |
|
612 |
|
613 section {* Pure types *} |
|
614 |
|
615 text {* Pure types will have always empty support. *} |
|
616 |
|
617 class pure = pt + |
|
618 assumes permute_pure: "p \<bullet> x = x" |
|
619 |
|
620 text {* Types @{typ unit} and @{typ bool} are pure. *} |
|
621 |
|
622 instance unit :: pure |
|
623 proof qed (rule permute_unit_def) |
|
624 |
|
625 instance bool :: pure |
|
626 proof qed (rule permute_bool_def) |
|
627 |
|
628 |
|
629 text {* Other type constructors preserve purity. *} |
|
630 |
|
631 instance "fun" :: (pure, pure) pure |
|
632 by default (simp add: permute_fun_def permute_pure) |
|
633 |
|
634 instance set :: (pure) pure |
|
635 by default (simp add: permute_set_def permute_pure) |
|
636 |
|
637 instance prod :: (pure, pure) pure |
|
638 by default (induct_tac x, simp add: permute_pure) |
|
639 |
|
640 instance sum :: (pure, pure) pure |
|
641 by default (induct_tac x, simp_all add: permute_pure) |
|
642 |
|
643 instance list :: (pure) pure |
|
644 by default (induct_tac x, simp_all add: permute_pure) |
|
645 |
|
646 instance option :: (pure) pure |
|
647 by default (induct_tac x, simp_all add: permute_pure) |
|
648 |
|
649 |
|
650 subsection {* Types @{typ char}, @{typ nat}, and @{typ int} *} |
|
651 |
|
652 instance char :: pure |
|
653 proof qed (rule permute_char_def) |
|
654 |
|
655 instance nat :: pure |
|
656 proof qed (rule permute_nat_def) |
|
657 |
|
658 instance int :: pure |
|
659 proof qed (rule permute_int_def) |
|
660 |
|
661 |
|
662 section {* Infrastructure for Equivariance and Perm_simp *} |
|
663 |
|
664 subsection {* Basic functions about permutations *} |
|
665 |
|
666 use "nominal_basics.ML" |
|
667 |
|
668 |
|
669 subsection {* Eqvt infrastructure *} |
|
670 |
|
671 text {* Setup of the theorem attributes @{text eqvt} and @{text eqvt_raw} *} |
|
672 |
|
673 use "nominal_thmdecls.ML" |
|
674 setup "Nominal_ThmDecls.setup" |
|
675 |
|
676 |
|
677 lemmas [eqvt] = |
|
678 (* pt types *) |
|
679 permute_prod.simps |
|
680 permute_list.simps |
|
681 permute_option.simps |
|
682 permute_sum.simps |
|
683 |
|
684 (* sets *) |
|
685 empty_eqvt insert_eqvt set_eqvt |
|
686 |
|
687 (* fsets *) |
|
688 permute_fset fset_eqvt |
|
689 |
|
690 (* multisets *) |
|
691 permute_multiset |
|
692 |
|
693 subsection {* perm_simp infrastructure *} |
|
694 |
|
695 definition |
|
696 "unpermute p = permute (- p)" |
|
697 |
|
698 lemma eqvt_apply: |
|
699 fixes f :: "'a::pt \<Rightarrow> 'b::pt" |
|
700 and x :: "'a::pt" |
|
701 shows "p \<bullet> (f x) \<equiv> (p \<bullet> f) (p \<bullet> x)" |
|
702 unfolding permute_fun_def by simp |
|
703 |
|
704 lemma eqvt_lambda: |
|
705 fixes f :: "'a::pt \<Rightarrow> 'b::pt" |
|
706 shows "p \<bullet> f \<equiv> (\<lambda>x. p \<bullet> (f (unpermute p x)))" |
|
707 unfolding permute_fun_def unpermute_def by simp |
|
708 |
|
709 lemma eqvt_bound: |
|
710 shows "p \<bullet> unpermute p x \<equiv> x" |
|
711 unfolding unpermute_def by simp |
|
712 |
|
713 text {* provides perm_simp methods *} |
|
714 |
|
715 use "nominal_permeq.ML" |
|
716 |
|
717 method_setup perm_simp = |
|
718 {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_simp_meth *} |
|
719 {* pushes permutations inside. *} |
|
720 |
|
721 method_setup perm_strict_simp = |
|
722 {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_strict_simp_meth *} |
|
723 {* pushes permutations inside, raises an error if it cannot solve all permutations. *} |
|
724 |
|
725 |
|
726 subsubsection {* Equivariance for permutations and swapping *} |
|
727 |
|
728 lemma permute_eqvt: |
|
729 shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)" |
|
730 unfolding permute_perm_def by simp |
|
731 |
|
732 (* the normal version of this lemma would cause loops *) |
|
733 lemma permute_eqvt_raw [eqvt_raw]: |
|
734 shows "p \<bullet> permute \<equiv> permute" |
|
735 apply(simp add: fun_eq_iff permute_fun_def) |
|
736 apply(subst permute_eqvt) |
|
737 apply(simp) |
|
738 done |
|
739 |
|
740 lemma zero_perm_eqvt [eqvt]: |
|
741 shows "p \<bullet> (0::perm) = 0" |
|
742 unfolding permute_perm_def by simp |
|
743 |
|
744 lemma add_perm_eqvt [eqvt]: |
|
745 fixes p p1 p2 :: perm |
|
746 shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2" |
|
747 unfolding permute_perm_def |
|
748 by (simp add: perm_eq_iff) |
|
749 |
|
750 lemma swap_eqvt [eqvt]: |
|
751 shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)" |
|
752 unfolding permute_perm_def |
|
753 by (auto simp add: swap_atom perm_eq_iff) |
|
754 |
|
755 lemma uminus_eqvt [eqvt]: |
|
756 fixes p q::"perm" |
|
757 shows "p \<bullet> (- q) = - (p \<bullet> q)" |
|
758 unfolding permute_perm_def |
|
759 by (simp add: diff_minus minus_add add_assoc) |
|
760 |
|
761 subsubsection {* Equivariance of Logical Operators *} |
|
762 |
|
763 lemma eq_eqvt [eqvt]: |
|
764 shows "p \<bullet> (x = y) \<longleftrightarrow> (p \<bullet> x) = (p \<bullet> y)" |
|
765 unfolding permute_eq_iff permute_bool_def .. |
|
766 |
|
767 lemma Not_eqvt [eqvt]: |
|
768 shows "p \<bullet> (\<not> A) \<longleftrightarrow> \<not> (p \<bullet> A)" |
|
769 by (simp add: permute_bool_def) |
|
770 |
|
771 lemma conj_eqvt [eqvt]: |
|
772 shows "p \<bullet> (A \<and> B) \<longleftrightarrow> (p \<bullet> A) \<and> (p \<bullet> B)" |
|
773 by (simp add: permute_bool_def) |
|
774 |
|
775 lemma imp_eqvt [eqvt]: |
|
776 shows "p \<bullet> (A \<longrightarrow> B) \<longleftrightarrow> (p \<bullet> A) \<longrightarrow> (p \<bullet> B)" |
|
777 by (simp add: permute_bool_def) |
|
778 |
|
779 declare imp_eqvt[folded induct_implies_def, eqvt] |
|
780 |
|
781 lemma all_eqvt [eqvt]: |
|
782 shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)" |
|
783 unfolding All_def |
|
784 by (perm_simp) (rule refl) |
|
785 |
|
786 declare all_eqvt[folded induct_forall_def, eqvt] |
|
787 |
|
788 lemma ex_eqvt [eqvt]: |
|
789 shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)" |
|
790 unfolding Ex_def |
|
791 by (perm_simp) (rule refl) |
|
792 |
|
793 lemma ex1_eqvt [eqvt]: |
|
794 shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)" |
|
795 unfolding Ex1_def |
|
796 by (perm_simp) (rule refl) |
|
797 |
|
798 lemma if_eqvt [eqvt]: |
|
799 shows "p \<bullet> (if b then x else y) = (if p \<bullet> b then p \<bullet> x else p \<bullet> y)" |
|
800 by (simp add: permute_fun_def permute_bool_def) |
|
801 |
|
802 lemma True_eqvt [eqvt]: |
|
803 shows "p \<bullet> True = True" |
|
804 unfolding permute_bool_def .. |
|
805 |
|
806 lemma False_eqvt [eqvt]: |
|
807 shows "p \<bullet> False = False" |
|
808 unfolding permute_bool_def .. |
|
809 |
|
810 lemma disj_eqvt [eqvt]: |
|
811 shows "p \<bullet> (A \<or> B) \<longleftrightarrow> (p \<bullet> A) \<or> (p \<bullet> B)" |
|
812 by (simp add: permute_bool_def) |
|
813 |
|
814 lemma all_eqvt2: |
|
815 shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. p \<bullet> P (- p \<bullet> x))" |
|
816 by (perm_simp add: permute_minus_cancel) (rule refl) |
|
817 |
|
818 lemma ex_eqvt2: |
|
819 shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. p \<bullet> P (- p \<bullet> x))" |
|
820 by (perm_simp add: permute_minus_cancel) (rule refl) |
|
821 |
|
822 lemma ex1_eqvt2: |
|
823 shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. p \<bullet> P (- p \<bullet> x))" |
|
824 by (perm_simp add: permute_minus_cancel) (rule refl) |
|
825 |
|
826 lemma the_eqvt: |
|
827 assumes unique: "\<exists>!x. P x" |
|
828 shows "(p \<bullet> (THE x. P x)) = (THE x. (p \<bullet> P) x)" |
|
829 apply(rule the1_equality [symmetric]) |
|
830 apply(rule_tac p="-p" in permute_boolE) |
|
831 apply(perm_simp add: permute_minus_cancel) |
|
832 apply(rule unique) |
|
833 apply(rule_tac p="-p" in permute_boolE) |
|
834 apply(perm_simp add: permute_minus_cancel) |
|
835 apply(rule theI'[OF unique]) |
|
836 done |
|
837 |
|
838 lemma the_eqvt2: |
|
839 assumes unique: "\<exists>!x. P x" |
|
840 shows "(p \<bullet> (THE x. P x)) = (THE x. p \<bullet> P (- p \<bullet> x))" |
|
841 apply(rule the1_equality [symmetric]) |
|
842 apply(simp add: ex1_eqvt2[symmetric]) |
|
843 apply(simp add: permute_bool_def unique) |
|
844 apply(simp add: permute_bool_def) |
|
845 apply(rule theI'[OF unique]) |
|
846 done |
|
847 |
|
848 subsubsection {* Equivariance of Set operators *} |
|
849 |
|
850 lemma mem_eqvt [eqvt]: |
|
851 shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)" |
|
852 unfolding permute_bool_def permute_set_def |
|
853 by (auto) |
|
854 |
|
855 lemma Collect_eqvt [eqvt]: |
|
856 shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}" |
|
857 unfolding permute_set_eq permute_fun_def |
|
858 by (auto simp add: permute_bool_def) |
|
859 |
|
860 lemma inter_eqvt [eqvt]: |
|
861 shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)" |
|
862 unfolding Int_def |
|
863 by (perm_simp) (rule refl) |
|
864 |
|
865 lemma Bex_eqvt [eqvt]: |
|
866 shows "p \<bullet> (\<exists>x \<in> S. P x) = (\<exists>x \<in> (p \<bullet> S). (p \<bullet> P) x)" |
|
867 unfolding Bex_def |
|
868 by (perm_simp) (rule refl) |
|
869 |
|
870 lemma Ball_eqvt [eqvt]: |
|
871 shows "p \<bullet> (\<forall>x \<in> S. P x) = (\<forall>x \<in> (p \<bullet> S). (p \<bullet> P) x)" |
|
872 unfolding Ball_def |
|
873 by (perm_simp) (rule refl) |
|
874 |
|
875 lemma image_eqvt [eqvt]: |
|
876 shows "p \<bullet> (f ` A) = (p \<bullet> f) ` (p \<bullet> A)" |
|
877 unfolding image_def |
|
878 by (perm_simp) (rule refl) |
|
879 |
|
880 lemma Image_eqvt [eqvt]: |
|
881 shows "p \<bullet> (R `` A) = (p \<bullet> R) `` (p \<bullet> A)" |
|
882 unfolding Image_def |
|
883 by (perm_simp) (rule refl) |
|
884 |
|
885 lemma UNIV_eqvt [eqvt]: |
|
886 shows "p \<bullet> UNIV = UNIV" |
|
887 unfolding UNIV_def |
|
888 by (perm_simp) (rule refl) |
|
889 |
|
890 lemma union_eqvt [eqvt]: |
|
891 shows "p \<bullet> (A \<union> B) = (p \<bullet> A) \<union> (p \<bullet> B)" |
|
892 unfolding Un_def |
|
893 by (perm_simp) (rule refl) |
|
894 |
|
895 lemma Diff_eqvt [eqvt]: |
|
896 fixes A B :: "'a::pt set" |
|
897 shows "p \<bullet> (A - B) = (p \<bullet> A) - (p \<bullet> B)" |
|
898 unfolding set_diff_eq |
|
899 by (perm_simp) (rule refl) |
|
900 |
|
901 lemma Compl_eqvt [eqvt]: |
|
902 fixes A :: "'a::pt set" |
|
903 shows "p \<bullet> (- A) = - (p \<bullet> A)" |
|
904 unfolding Compl_eq_Diff_UNIV |
|
905 by (perm_simp) (rule refl) |
|
906 |
|
907 lemma subset_eqvt [eqvt]: |
|
908 shows "p \<bullet> (S \<subseteq> T) \<longleftrightarrow> (p \<bullet> S) \<subseteq> (p \<bullet> T)" |
|
909 unfolding subset_eq |
|
910 by (perm_simp) (rule refl) |
|
911 |
|
912 lemma psubset_eqvt [eqvt]: |
|
913 shows "p \<bullet> (S \<subset> T) \<longleftrightarrow> (p \<bullet> S) \<subset> (p \<bullet> T)" |
|
914 unfolding psubset_eq |
|
915 by (perm_simp) (rule refl) |
|
916 |
|
917 lemma vimage_eqvt [eqvt]: |
|
918 shows "p \<bullet> (f -` A) = (p \<bullet> f) -` (p \<bullet> A)" |
|
919 unfolding vimage_def |
|
920 by (perm_simp) (rule refl) |
|
921 |
|
922 lemma Union_eqvt [eqvt]: |
|
923 shows "p \<bullet> (\<Union> S) = \<Union> (p \<bullet> S)" |
|
924 unfolding Union_eq |
|
925 by (perm_simp) (rule refl) |
|
926 |
|
927 lemma Inter_eqvt [eqvt]: |
|
928 shows "p \<bullet> (\<Inter> S) = \<Inter> (p \<bullet> S)" |
|
929 unfolding Inter_eq |
|
930 by (perm_simp) (rule refl) |
|
931 |
|
932 (* FIXME: eqvt attribute *) |
|
933 lemma Sigma_eqvt: |
|
934 shows "(p \<bullet> (X \<times> Y)) = (p \<bullet> X) \<times> (p \<bullet> Y)" |
|
935 unfolding Sigma_def |
|
936 unfolding SUP_def |
|
937 by (perm_simp) (rule refl) |
|
938 |
|
939 text {* |
|
940 In order to prove that lfp is equivariant we need two |
|
941 auxiliary classes which specify that (op <=) and |
|
942 Inf are equivariant. Instances for bool and fun are |
|
943 given. |
|
944 *} |
|
945 |
|
946 class le_eqvt = order + |
|
947 assumes le_eqvt [eqvt]: "p \<bullet> (x \<le> y) = ((p \<bullet> x) \<le> (p \<bullet> (y::('a::{pt, order}))))" |
|
948 |
|
949 class inf_eqvt = complete_lattice + |
|
950 assumes inf_eqvt [eqvt]: "p \<bullet> (Inf X) = Inf (p \<bullet> (X::('a::{pt, Inf}) set))" |
|
951 |
|
952 instantiation bool :: le_eqvt |
|
953 begin |
|
954 |
|
955 instance |
|
956 apply(default) |
|
957 apply perm_simp |
|
958 apply(rule refl) |
|
959 done |
|
960 |
|
961 end |
|
962 |
|
963 instantiation "fun" :: (pt, le_eqvt) le_eqvt |
|
964 begin |
|
965 |
|
966 instance |
|
967 apply(default) |
|
968 unfolding le_fun_def |
|
969 apply(perm_simp) |
|
970 apply(rule refl) |
|
971 done |
|
972 |
|
973 end |
|
974 |
|
975 instantiation bool :: inf_eqvt |
|
976 begin |
|
977 |
|
978 instance |
|
979 apply(default) |
|
980 apply(perm_simp) |
|
981 apply(rule refl) |
|
982 done |
|
983 |
|
984 end |
|
985 |
|
986 instantiation "fun" :: (pt, inf_eqvt) inf_eqvt |
|
987 begin |
|
988 |
|
989 instance |
|
990 apply(default) |
|
991 unfolding Inf_fun_def INF_def |
|
992 apply(perm_simp) |
|
993 apply(rule refl) |
|
994 done |
|
995 |
|
996 end |
|
997 |
|
998 lemma lfp_eqvt [eqvt]: |
|
999 fixes F::"('a \<Rightarrow> 'b) \<Rightarrow> ('a::pt \<Rightarrow> 'b::{inf_eqvt, le_eqvt})" |
|
1000 shows "p \<bullet> (lfp F) = lfp (p \<bullet> F)" |
|
1001 unfolding lfp_def |
|
1002 by (perm_simp) (rule refl) |
|
1003 |
|
1004 lemma finite_eqvt [eqvt]: |
|
1005 shows "p \<bullet> finite A = finite (p \<bullet> A)" |
|
1006 unfolding finite_def |
|
1007 by (perm_simp) (rule refl) |
|
1008 |
|
1009 |
|
1010 subsubsection {* Equivariance for product operations *} |
|
1011 |
|
1012 lemma fst_eqvt [eqvt]: |
|
1013 shows "p \<bullet> (fst x) = fst (p \<bullet> x)" |
|
1014 by (cases x) simp |
|
1015 |
|
1016 lemma snd_eqvt [eqvt]: |
|
1017 shows "p \<bullet> (snd x) = snd (p \<bullet> x)" |
|
1018 by (cases x) simp |
|
1019 |
|
1020 lemma split_eqvt [eqvt]: |
|
1021 shows "p \<bullet> (split P x) = split (p \<bullet> P) (p \<bullet> x)" |
|
1022 unfolding split_def |
|
1023 by (perm_simp) (rule refl) |
|
1024 |
|
1025 |
|
1026 subsubsection {* Equivariance for list operations *} |
|
1027 |
|
1028 lemma append_eqvt [eqvt]: |
|
1029 shows "p \<bullet> (xs @ ys) = (p \<bullet> xs) @ (p \<bullet> ys)" |
|
1030 by (induct xs) auto |
|
1031 |
|
1032 lemma rev_eqvt [eqvt]: |
|
1033 shows "p \<bullet> (rev xs) = rev (p \<bullet> xs)" |
|
1034 by (induct xs) (simp_all add: append_eqvt) |
|
1035 |
|
1036 lemma map_eqvt [eqvt]: |
|
1037 shows "p \<bullet> (map f xs) = map (p \<bullet> f) (p \<bullet> xs)" |
|
1038 by (induct xs) (simp_all, simp only: permute_fun_app_eq) |
|
1039 |
|
1040 lemma removeAll_eqvt [eqvt]: |
|
1041 shows "p \<bullet> (removeAll x xs) = removeAll (p \<bullet> x) (p \<bullet> xs)" |
|
1042 by (induct xs) (auto) |
|
1043 |
|
1044 lemma filter_eqvt [eqvt]: |
|
1045 shows "p \<bullet> (filter f xs) = filter (p \<bullet> f) (p \<bullet> xs)" |
|
1046 apply(induct xs) |
|
1047 apply(simp) |
|
1048 apply(simp only: filter.simps permute_list.simps if_eqvt) |
|
1049 apply(simp only: permute_fun_app_eq) |
|
1050 done |
|
1051 |
|
1052 lemma distinct_eqvt [eqvt]: |
|
1053 shows "p \<bullet> (distinct xs) = distinct (p \<bullet> xs)" |
|
1054 apply(induct xs) |
|
1055 apply(simp add: permute_bool_def) |
|
1056 apply(simp add: conj_eqvt Not_eqvt mem_eqvt set_eqvt) |
|
1057 done |
|
1058 |
|
1059 lemma length_eqvt [eqvt]: |
|
1060 shows "p \<bullet> (length xs) = length (p \<bullet> xs)" |
|
1061 by (induct xs) (simp_all add: permute_pure) |
|
1062 |
|
1063 |
|
1064 subsubsection {* Equivariance for @{typ "'a option"} *} |
|
1065 |
|
1066 lemma option_map_eqvt[eqvt]: |
|
1067 shows "p \<bullet> (Option.map f x) = Option.map (p \<bullet> f) (p \<bullet> x)" |
|
1068 by (cases x) (simp_all, simp add: permute_fun_app_eq) |
|
1069 |
|
1070 |
|
1071 subsubsection {* Equivariance for @{typ "'a fset"} *} |
|
1072 |
|
1073 lemma in_fset_eqvt [eqvt]: |
|
1074 shows "(p \<bullet> (x |\<in>| S)) = ((p \<bullet> x) |\<in>| (p \<bullet> S))" |
|
1075 unfolding in_fset |
|
1076 by (perm_simp) (simp) |
|
1077 |
|
1078 lemma union_fset_eqvt [eqvt]: |
|
1079 shows "(p \<bullet> (S |\<union>| T)) = ((p \<bullet> S) |\<union>| (p \<bullet> T))" |
|
1080 by (induct S) (simp_all) |
|
1081 |
|
1082 lemma inter_list_eqvt [eqvt]: |
|
1083 shows "p \<bullet> (inter_list S T) = inter_list (p \<bullet> S) (p \<bullet> T)" |
|
1084 unfolding list_eq_def inter_list_def |
|
1085 by perm_simp simp |
|
1086 |
|
1087 lemma inter_fset_eqvt [eqvt]: |
|
1088 shows "(p \<bullet> (S |\<inter>| T)) = ((p \<bullet> S) |\<inter>| (p \<bullet> T))" |
|
1089 by (lifting inter_list_eqvt) |
|
1090 |
|
1091 lemma sub_list_eqvt [eqvt]: |
|
1092 shows "p \<bullet> (sub_list S T) = sub_list (p \<bullet> S) (p \<bullet> T)" |
|
1093 unfolding sub_list_def |
|
1094 by perm_simp simp |
|
1095 |
|
1096 lemma subset_fset_eqvt [eqvt]: |
|
1097 shows "(p \<bullet> (S |\<subseteq>| T)) = ((p \<bullet> S) |\<subseteq>| (p \<bullet> T))" |
|
1098 by (lifting sub_list_eqvt) |
|
1099 |
|
1100 lemma map_fset_eqvt [eqvt]: |
|
1101 shows "p \<bullet> (map_fset f S) = map_fset (p \<bullet> f) (p \<bullet> S)" |
|
1102 by (lifting map_eqvt) |
|
1103 |
|
1104 |
|
1105 section {* Supp, Freshness and Supports *} |
|
1106 |
|
1107 context pt |
|
1108 begin |
|
1109 |
|
1110 definition |
|
1111 supp :: "'a \<Rightarrow> atom set" |
|
1112 where |
|
1113 "supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}" |
|
1114 |
|
1115 definition |
|
1116 fresh :: "atom \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55) |
|
1117 where |
|
1118 "a \<sharp> x \<equiv> a \<notin> supp x" |
|
1119 |
|
1120 end |
|
1121 |
|
1122 lemma supp_conv_fresh: |
|
1123 shows "supp x = {a. \<not> a \<sharp> x}" |
|
1124 unfolding fresh_def by simp |
|
1125 |
|
1126 lemma swap_rel_trans: |
|
1127 assumes "sort_of a = sort_of b" |
|
1128 assumes "sort_of b = sort_of c" |
|
1129 assumes "(a \<rightleftharpoons> c) \<bullet> x = x" |
|
1130 assumes "(b \<rightleftharpoons> c) \<bullet> x = x" |
|
1131 shows "(a \<rightleftharpoons> b) \<bullet> x = x" |
|
1132 proof (cases) |
|
1133 assume "a = b \<or> c = b" |
|
1134 with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto |
|
1135 next |
|
1136 assume *: "\<not> (a = b \<or> c = b)" |
|
1137 have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x" |
|
1138 using assms by simp |
|
1139 also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)" |
|
1140 using assms * by (simp add: swap_triple) |
|
1141 finally show "(a \<rightleftharpoons> b) \<bullet> x = x" . |
|
1142 qed |
|
1143 |
|
1144 lemma swap_fresh_fresh: |
|
1145 assumes a: "a \<sharp> x" |
|
1146 and b: "b \<sharp> x" |
|
1147 shows "(a \<rightleftharpoons> b) \<bullet> x = x" |
|
1148 proof (cases) |
|
1149 assume asm: "sort_of a = sort_of b" |
|
1150 have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}" |
|
1151 using a b unfolding fresh_def supp_def by simp_all |
|
1152 then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp |
|
1153 then obtain c |
|
1154 where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b" |
|
1155 by (rule obtain_atom) (auto) |
|
1156 then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all) |
|
1157 next |
|
1158 assume "sort_of a \<noteq> sort_of b" |
|
1159 then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp |
|
1160 qed |
|
1161 |
|
1162 |
|
1163 subsection {* supp and fresh are equivariant *} |
|
1164 |
|
1165 |
|
1166 lemma supp_eqvt [eqvt]: |
|
1167 shows "p \<bullet> (supp x) = supp (p \<bullet> x)" |
|
1168 unfolding supp_def |
|
1169 by (perm_simp) |
|
1170 (simp only: permute_eqvt[symmetric]) |
|
1171 |
|
1172 lemma fresh_eqvt [eqvt]: |
|
1173 shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)" |
|
1174 unfolding fresh_def |
|
1175 by (perm_simp) (rule refl) |
|
1176 |
|
1177 lemma fresh_permute_iff: |
|
1178 shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x" |
|
1179 by (simp only: fresh_eqvt[symmetric] permute_bool_def) |
|
1180 |
|
1181 lemma fresh_permute_left: |
|
1182 shows "a \<sharp> p \<bullet> x \<longleftrightarrow> - p \<bullet> a \<sharp> x" |
|
1183 proof |
|
1184 assume "a \<sharp> p \<bullet> x" |
|
1185 then have "- p \<bullet> a \<sharp> - p \<bullet> p \<bullet> x" by (simp only: fresh_permute_iff) |
|
1186 then show "- p \<bullet> a \<sharp> x" by simp |
|
1187 next |
|
1188 assume "- p \<bullet> a \<sharp> x" |
|
1189 then have "p \<bullet> - p \<bullet> a \<sharp> p \<bullet> x" by (simp only: fresh_permute_iff) |
|
1190 then show "a \<sharp> p \<bullet> x" by simp |
|
1191 qed |
|
1192 |
|
1193 |
|
1194 section {* supports *} |
|
1195 |
|
1196 definition |
|
1197 supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80) |
|
1198 where |
|
1199 "S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)" |
|
1200 |
|
1201 lemma supp_is_subset: |
|
1202 fixes S :: "atom set" |
|
1203 and x :: "'a::pt" |
|
1204 assumes a1: "S supports x" |
|
1205 and a2: "finite S" |
|
1206 shows "(supp x) \<subseteq> S" |
|
1207 proof (rule ccontr) |
|
1208 assume "\<not> (supp x \<subseteq> S)" |
|
1209 then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto |
|
1210 from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" unfolding supports_def by auto |
|
1211 then have "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto |
|
1212 with a2 have "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}" by (simp add: finite_subset) |
|
1213 then have "a \<notin> (supp x)" unfolding supp_def by simp |
|
1214 with b1 show False by simp |
|
1215 qed |
|
1216 |
|
1217 lemma supports_finite: |
|
1218 fixes S :: "atom set" |
|
1219 and x :: "'a::pt" |
|
1220 assumes a1: "S supports x" |
|
1221 and a2: "finite S" |
|
1222 shows "finite (supp x)" |
|
1223 proof - |
|
1224 have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset) |
|
1225 then show "finite (supp x)" using a2 by (simp add: finite_subset) |
|
1226 qed |
|
1227 |
|
1228 lemma supp_supports: |
|
1229 fixes x :: "'a::pt" |
|
1230 shows "(supp x) supports x" |
|
1231 unfolding supports_def |
|
1232 proof (intro strip) |
|
1233 fix a b |
|
1234 assume "a \<notin> (supp x) \<and> b \<notin> (supp x)" |
|
1235 then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def) |
|
1236 then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh) |
|
1237 qed |
|
1238 |
|
1239 lemma supports_fresh: |
|
1240 fixes x :: "'a::pt" |
|
1241 assumes a1: "S supports x" |
|
1242 and a2: "finite S" |
|
1243 and a3: "a \<notin> S" |
|
1244 shows "a \<sharp> x" |
|
1245 unfolding fresh_def |
|
1246 proof - |
|
1247 have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset) |
|
1248 then show "a \<notin> (supp x)" using a3 by auto |
|
1249 qed |
|
1250 |
|
1251 lemma supp_is_least_supports: |
|
1252 fixes S :: "atom set" |
|
1253 and x :: "'a::pt" |
|
1254 assumes a1: "S supports x" |
|
1255 and a2: "finite S" |
|
1256 and a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'" |
|
1257 shows "(supp x) = S" |
|
1258 proof (rule equalityI) |
|
1259 show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset) |
|
1260 with a2 have fin: "finite (supp x)" by (rule rev_finite_subset) |
|
1261 have "(supp x) supports x" by (rule supp_supports) |
|
1262 with fin a3 show "S \<subseteq> supp x" by blast |
|
1263 qed |
|
1264 |
|
1265 |
|
1266 lemma subsetCI: |
|
1267 shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B" |
|
1268 by auto |
|
1269 |
|
1270 lemma finite_supp_unique: |
|
1271 assumes a1: "S supports x" |
|
1272 assumes a2: "finite S" |
|
1273 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" |
|
1274 shows "(supp x) = S" |
|
1275 using a1 a2 |
|
1276 proof (rule supp_is_least_supports) |
|
1277 fix S' |
|
1278 assume "finite S'" and "S' supports x" |
|
1279 show "S \<subseteq> S'" |
|
1280 proof (rule subsetCI) |
|
1281 fix a |
|
1282 assume "a \<in> S" and "a \<notin> S'" |
|
1283 have "finite (S \<union> S')" |
|
1284 using `finite S` `finite S'` by simp |
|
1285 then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a" |
|
1286 by (rule obtain_atom) |
|
1287 then have "b \<notin> S" and "b \<notin> S'" and "sort_of a = sort_of b" |
|
1288 by simp_all |
|
1289 then have "(a \<rightleftharpoons> b) \<bullet> x = x" |
|
1290 using `a \<notin> S'` `S' supports x` by (simp add: supports_def) |
|
1291 moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x" |
|
1292 using `a \<in> S` `b \<notin> S` `sort_of a = sort_of b` |
|
1293 by (rule a3) |
|
1294 ultimately show "False" by simp |
|
1295 qed |
|
1296 qed |
|
1297 |
|
1298 section {* Support w.r.t. relations *} |
|
1299 |
|
1300 text {* |
|
1301 This definition is used for unquotient types, where |
|
1302 alpha-equivalence does not coincide with equality. |
|
1303 *} |
|
1304 |
|
1305 definition |
|
1306 "supp_rel R x = {a. infinite {b. \<not>(R ((a \<rightleftharpoons> b) \<bullet> x) x)}}" |
|
1307 |
|
1308 |
|
1309 |
|
1310 section {* Finitely-supported types *} |
|
1311 |
|
1312 class fs = pt + |
|
1313 assumes finite_supp: "finite (supp x)" |
|
1314 |
|
1315 lemma pure_supp: |
|
1316 fixes x::"'a::pure" |
|
1317 shows "supp x = {}" |
|
1318 unfolding supp_def by (simp add: permute_pure) |
|
1319 |
|
1320 lemma pure_fresh: |
|
1321 fixes x::"'a::pure" |
|
1322 shows "a \<sharp> x" |
|
1323 unfolding fresh_def by (simp add: pure_supp) |
|
1324 |
|
1325 instance pure < fs |
|
1326 by default (simp add: pure_supp) |
|
1327 |
|
1328 |
|
1329 subsection {* Type @{typ atom} is finitely-supported. *} |
|
1330 |
|
1331 lemma supp_atom: |
|
1332 shows "supp a = {a}" |
|
1333 by (rule finite_supp_unique) |
|
1334 (auto simp add: supports_def) |
|
1335 |
|
1336 lemma fresh_atom: |
|
1337 shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b" |
|
1338 unfolding fresh_def supp_atom by simp |
|
1339 |
|
1340 instance atom :: fs |
|
1341 by default (simp add: supp_atom) |
|
1342 |
|
1343 |
|
1344 section {* Type @{typ perm} is finitely-supported. *} |
|
1345 |
|
1346 lemma perm_swap_eq: |
|
1347 shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)" |
|
1348 unfolding permute_perm_def |
|
1349 by (metis add_diff_cancel minus_perm_def) |
|
1350 |
|
1351 lemma supports_perm: |
|
1352 shows "{a. p \<bullet> a \<noteq> a} supports p" |
|
1353 unfolding supports_def |
|
1354 unfolding perm_swap_eq |
|
1355 by (simp add: swap_eqvt) |
|
1356 |
|
1357 lemma finite_perm_lemma: |
|
1358 shows "finite {a::atom. p \<bullet> a \<noteq> a}" |
|
1359 unfolding permute_atom_def |
|
1360 using finite_gpermute_neq . |
|
1361 |
|
1362 lemma supp_perm: |
|
1363 shows "supp p = {a. p \<bullet> a \<noteq> a}" |
|
1364 apply (rule finite_supp_unique) |
|
1365 apply (simp_all add: perm_swap_eq swap_eqvt supports_perm finite_perm_lemma) |
|
1366 apply (auto simp add: perm_eq_iff swap_atom perm_swap_eq swap_eqvt) |
|
1367 done |
|
1368 |
|
1369 lemma fresh_perm: |
|
1370 shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a" |
|
1371 unfolding fresh_def |
|
1372 by (simp add: supp_perm) |
|
1373 |
|
1374 lemma supp_swap: |
|
1375 shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})" |
|
1376 by (auto simp add: supp_perm swap_atom) |
|
1377 |
|
1378 lemma fresh_zero_perm: |
|
1379 shows "a \<sharp> (0::perm)" |
|
1380 unfolding fresh_perm by simp |
|
1381 |
|
1382 lemma supp_zero_perm: |
|
1383 shows "supp (0::perm) = {}" |
|
1384 unfolding supp_perm by simp |
|
1385 |
|
1386 lemma fresh_plus_perm: |
|
1387 fixes p q::perm |
|
1388 assumes "a \<sharp> p" "a \<sharp> q" |
|
1389 shows "a \<sharp> (p + q)" |
|
1390 using assms |
|
1391 unfolding fresh_def |
|
1392 by (auto simp add: supp_perm) |
|
1393 |
|
1394 lemma supp_plus_perm: |
|
1395 fixes p q::perm |
|
1396 shows "supp (p + q) \<subseteq> supp p \<union> supp q" |
|
1397 by (auto simp add: supp_perm) |
|
1398 |
|
1399 lemma fresh_minus_perm: |
|
1400 fixes p::perm |
|
1401 shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p" |
|
1402 unfolding fresh_def supp_perm |
|
1403 by (simp) (metis permute_minus_cancel(1)) |
|
1404 |
|
1405 lemma supp_minus_perm: |
|
1406 fixes p::perm |
|
1407 shows "supp (- p) = supp p" |
|
1408 unfolding supp_conv_fresh |
|
1409 by (simp add: fresh_minus_perm) |
|
1410 |
|
1411 lemma plus_perm_eq: |
|
1412 fixes p q::"perm" |
|
1413 assumes asm: "supp p \<inter> supp q = {}" |
|
1414 shows "p + q = q + p" |
|
1415 unfolding perm_eq_iff |
|
1416 proof |
|
1417 fix a::"atom" |
|
1418 show "(p + q) \<bullet> a = (q + p) \<bullet> a" |
|
1419 proof - |
|
1420 { assume "a \<notin> supp p" "a \<notin> supp q" |
|
1421 then have "(p + q) \<bullet> a = (q + p) \<bullet> a" |
|
1422 by (simp add: supp_perm) |
|
1423 } |
|
1424 moreover |
|
1425 { assume a: "a \<in> supp p" "a \<notin> supp q" |
|
1426 then have "p \<bullet> a \<in> supp p" by (simp add: supp_perm) |
|
1427 then have "p \<bullet> a \<notin> supp q" using asm by auto |
|
1428 with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" |
|
1429 by (simp add: supp_perm) |
|
1430 } |
|
1431 moreover |
|
1432 { assume a: "a \<notin> supp p" "a \<in> supp q" |
|
1433 then have "q \<bullet> a \<in> supp q" by (simp add: supp_perm) |
|
1434 then have "q \<bullet> a \<notin> supp p" using asm by auto |
|
1435 with a have "(p + q) \<bullet> a = (q + p) \<bullet> a" |
|
1436 by (simp add: supp_perm) |
|
1437 } |
|
1438 ultimately show "(p + q) \<bullet> a = (q + p) \<bullet> a" |
|
1439 using asm by blast |
|
1440 qed |
|
1441 qed |
|
1442 |
|
1443 lemma supp_plus_perm_eq: |
|
1444 fixes p q::perm |
|
1445 assumes asm: "supp p \<inter> supp q = {}" |
|
1446 shows "supp (p + q) = supp p \<union> supp q" |
|
1447 proof - |
|
1448 { fix a::"atom" |
|
1449 assume "a \<in> supp p" |
|
1450 then have "a \<notin> supp q" using asm by auto |
|
1451 then have "a \<in> supp (p + q)" using `a \<in> supp p` |
|
1452 by (simp add: supp_perm) |
|
1453 } |
|
1454 moreover |
|
1455 { fix a::"atom" |
|
1456 assume "a \<in> supp q" |
|
1457 then have "a \<notin> supp p" using asm by auto |
|
1458 then have "a \<in> supp (q + p)" using `a \<in> supp q` |
|
1459 by (simp add: supp_perm) |
|
1460 then have "a \<in> supp (p + q)" using asm plus_perm_eq |
|
1461 by metis |
|
1462 } |
|
1463 ultimately have "supp p \<union> supp q \<subseteq> supp (p + q)" |
|
1464 by blast |
|
1465 then show "supp (p + q) = supp p \<union> supp q" using supp_plus_perm |
|
1466 by blast |
|
1467 qed |
|
1468 |
|
1469 instance perm :: fs |
|
1470 by default (simp add: supp_perm finite_perm_lemma) |
|
1471 |
|
1472 |
|
1473 |
|
1474 section {* Finite Support instances for other types *} |
|
1475 |
|
1476 |
|
1477 subsection {* Type @{typ "'a \<times> 'b"} is finitely-supported. *} |
|
1478 |
|
1479 lemma supp_Pair: |
|
1480 shows "supp (x, y) = supp x \<union> supp y" |
|
1481 by (simp add: supp_def Collect_imp_eq Collect_neg_eq) |
|
1482 |
|
1483 lemma fresh_Pair: |
|
1484 shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y" |
|
1485 by (simp add: fresh_def supp_Pair) |
|
1486 |
|
1487 lemma supp_Unit: |
|
1488 shows "supp () = {}" |
|
1489 by (simp add: supp_def) |
|
1490 |
|
1491 lemma fresh_Unit: |
|
1492 shows "a \<sharp> ()" |
|
1493 by (simp add: fresh_def supp_Unit) |
|
1494 |
|
1495 instance prod :: (fs, fs) fs |
|
1496 by default (auto simp add: supp_Pair finite_supp) |
|
1497 |
|
1498 |
|
1499 subsection {* Type @{typ "'a + 'b"} is finitely supported *} |
|
1500 |
|
1501 lemma supp_Inl: |
|
1502 shows "supp (Inl x) = supp x" |
|
1503 by (simp add: supp_def) |
|
1504 |
|
1505 lemma supp_Inr: |
|
1506 shows "supp (Inr x) = supp x" |
|
1507 by (simp add: supp_def) |
|
1508 |
|
1509 lemma fresh_Inl: |
|
1510 shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x" |
|
1511 by (simp add: fresh_def supp_Inl) |
|
1512 |
|
1513 lemma fresh_Inr: |
|
1514 shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y" |
|
1515 by (simp add: fresh_def supp_Inr) |
|
1516 |
|
1517 instance sum :: (fs, fs) fs |
|
1518 apply default |
|
1519 apply (case_tac x) |
|
1520 apply (simp_all add: supp_Inl supp_Inr finite_supp) |
|
1521 done |
|
1522 |
|
1523 |
|
1524 subsection {* Type @{typ "'a option"} is finitely supported *} |
|
1525 |
|
1526 lemma supp_None: |
|
1527 shows "supp None = {}" |
|
1528 by (simp add: supp_def) |
|
1529 |
|
1530 lemma supp_Some: |
|
1531 shows "supp (Some x) = supp x" |
|
1532 by (simp add: supp_def) |
|
1533 |
|
1534 lemma fresh_None: |
|
1535 shows "a \<sharp> None" |
|
1536 by (simp add: fresh_def supp_None) |
|
1537 |
|
1538 lemma fresh_Some: |
|
1539 shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x" |
|
1540 by (simp add: fresh_def supp_Some) |
|
1541 |
|
1542 instance option :: (fs) fs |
|
1543 apply default |
|
1544 apply (induct_tac x) |
|
1545 apply (simp_all add: supp_None supp_Some finite_supp) |
|
1546 done |
|
1547 |
|
1548 |
|
1549 subsubsection {* Type @{typ "'a list"} is finitely supported *} |
|
1550 |
|
1551 lemma supp_Nil: |
|
1552 shows "supp [] = {}" |
|
1553 by (simp add: supp_def) |
|
1554 |
|
1555 lemma fresh_Nil: |
|
1556 shows "a \<sharp> []" |
|
1557 by (simp add: fresh_def supp_Nil) |
|
1558 |
|
1559 lemma supp_Cons: |
|
1560 shows "supp (x # xs) = supp x \<union> supp xs" |
|
1561 by (simp add: supp_def Collect_imp_eq Collect_neg_eq) |
|
1562 |
|
1563 lemma fresh_Cons: |
|
1564 shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs" |
|
1565 by (simp add: fresh_def supp_Cons) |
|
1566 |
|
1567 lemma supp_append: |
|
1568 shows "supp (xs @ ys) = supp xs \<union> supp ys" |
|
1569 by (induct xs) (auto simp add: supp_Nil supp_Cons) |
|
1570 |
|
1571 lemma fresh_append: |
|
1572 shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys" |
|
1573 by (induct xs) (simp_all add: fresh_Nil fresh_Cons) |
|
1574 |
|
1575 lemma supp_rev: |
|
1576 shows "supp (rev xs) = supp xs" |
|
1577 by (induct xs) (auto simp add: supp_append supp_Cons supp_Nil) |
|
1578 |
|
1579 lemma fresh_rev: |
|
1580 shows "a \<sharp> rev xs \<longleftrightarrow> a \<sharp> xs" |
|
1581 by (induct xs) (auto simp add: fresh_append fresh_Cons fresh_Nil) |
|
1582 |
|
1583 lemma supp_removeAll: |
|
1584 fixes x::"atom" |
|
1585 shows "supp (removeAll x xs) = supp xs - {x}" |
|
1586 by (induct xs) |
|
1587 (auto simp add: supp_Nil supp_Cons supp_atom) |
|
1588 |
|
1589 lemma supp_of_atom_list: |
|
1590 fixes as::"atom list" |
|
1591 shows "supp as = set as" |
|
1592 by (induct as) |
|
1593 (simp_all add: supp_Nil supp_Cons supp_atom) |
|
1594 |
|
1595 instance list :: (fs) fs |
|
1596 apply default |
|
1597 apply (induct_tac x) |
|
1598 apply (simp_all add: supp_Nil supp_Cons finite_supp) |
|
1599 done |
|
1600 |
|
1601 |
|
1602 section {* Support and Freshness for Applications *} |
|
1603 |
|
1604 lemma fresh_conv_MOST: |
|
1605 shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)" |
|
1606 unfolding fresh_def supp_def |
|
1607 unfolding MOST_iff_cofinite by simp |
|
1608 |
|
1609 lemma fresh_fun_app: |
|
1610 assumes "a \<sharp> f" and "a \<sharp> x" |
|
1611 shows "a \<sharp> f x" |
|
1612 using assms |
|
1613 unfolding fresh_conv_MOST |
|
1614 unfolding permute_fun_app_eq |
|
1615 by (elim MOST_rev_mp) (simp) |
|
1616 |
|
1617 lemma supp_fun_app: |
|
1618 shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)" |
|
1619 using fresh_fun_app |
|
1620 unfolding fresh_def |
|
1621 by auto |
|
1622 |
|
1623 |
|
1624 subsection {* Equivariance Predicate @{text eqvt} and @{text eqvt_at}*} |
|
1625 |
|
1626 definition |
|
1627 "eqvt f \<equiv> \<forall>p. p \<bullet> f = f" |
|
1628 |
|
1629 lemma eqvt_boolI: |
|
1630 fixes f::"bool" |
|
1631 shows "eqvt f" |
|
1632 unfolding eqvt_def |
|
1633 by (simp add: permute_bool_def) |
|
1634 |
|
1635 |
|
1636 text {* equivariance of a function at a given argument *} |
|
1637 |
|
1638 definition |
|
1639 "eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)" |
|
1640 |
|
1641 lemma eqvtI: |
|
1642 shows "(\<And>p. p \<bullet> f \<equiv> f) \<Longrightarrow> eqvt f" |
|
1643 unfolding eqvt_def |
|
1644 by simp |
|
1645 |
|
1646 lemma eqvt_at_perm: |
|
1647 assumes "eqvt_at f x" |
|
1648 shows "eqvt_at f (q \<bullet> x)" |
|
1649 proof - |
|
1650 { fix p::"perm" |
|
1651 have "p \<bullet> (f (q \<bullet> x)) = p \<bullet> q \<bullet> (f x)" |
|
1652 using assms by (simp add: eqvt_at_def) |
|
1653 also have "\<dots> = (p + q) \<bullet> (f x)" by simp |
|
1654 also have "\<dots> = f ((p + q) \<bullet> x)" |
|
1655 using assms by (simp add: eqvt_at_def) |
|
1656 finally have "p \<bullet> (f (q \<bullet> x)) = f (p \<bullet> q \<bullet> x)" by simp } |
|
1657 then show "eqvt_at f (q \<bullet> x)" unfolding eqvt_at_def |
|
1658 by simp |
|
1659 qed |
|
1660 |
|
1661 lemma supp_fun_eqvt: |
|
1662 assumes a: "eqvt f" |
|
1663 shows "supp f = {}" |
|
1664 using a |
|
1665 unfolding eqvt_def |
|
1666 unfolding supp_def |
|
1667 by simp |
|
1668 |
|
1669 lemma fresh_fun_eqvt_app: |
|
1670 assumes a: "eqvt f" |
|
1671 shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x" |
|
1672 proof - |
|
1673 from a have "supp f = {}" by (simp add: supp_fun_eqvt) |
|
1674 then show "a \<sharp> x \<Longrightarrow> a \<sharp> f x" |
|
1675 unfolding fresh_def |
|
1676 using supp_fun_app by auto |
|
1677 qed |
|
1678 |
|
1679 lemma supp_fun_app_eqvt: |
|
1680 assumes a: "eqvt f" |
|
1681 shows "supp (f x) \<subseteq> supp x" |
|
1682 using fresh_fun_eqvt_app[OF a] |
|
1683 unfolding fresh_def |
|
1684 by auto |
|
1685 |
|
1686 lemma supp_eqvt_at: |
|
1687 assumes asm: "eqvt_at f x" |
|
1688 and fin: "finite (supp x)" |
|
1689 shows "supp (f x) \<subseteq> supp x" |
|
1690 apply(rule supp_is_subset) |
|
1691 unfolding supports_def |
|
1692 unfolding fresh_def[symmetric] |
|
1693 using asm |
|
1694 apply(simp add: eqvt_at_def swap_fresh_fresh) |
|
1695 apply(rule fin) |
|
1696 done |
|
1697 |
|
1698 lemma finite_supp_eqvt_at: |
|
1699 assumes asm: "eqvt_at f x" |
|
1700 and fin: "finite (supp x)" |
|
1701 shows "finite (supp (f x))" |
|
1702 apply(rule finite_subset) |
|
1703 apply(rule supp_eqvt_at[OF asm fin]) |
|
1704 apply(rule fin) |
|
1705 done |
|
1706 |
|
1707 lemma fresh_eqvt_at: |
|
1708 assumes asm: "eqvt_at f x" |
|
1709 and fin: "finite (supp x)" |
|
1710 and fresh: "a \<sharp> x" |
|
1711 shows "a \<sharp> f x" |
|
1712 using fresh |
|
1713 unfolding fresh_def |
|
1714 using supp_eqvt_at[OF asm fin] |
|
1715 by auto |
|
1716 |
|
1717 |
|
1718 subsection {* helper functions for nominal_functions *} |
|
1719 |
|
1720 lemma THE_defaultI2: |
|
1721 assumes "\<exists>!x. P x" "\<And>x. P x \<Longrightarrow> Q x" |
|
1722 shows "Q (THE_default d P)" |
|
1723 by (iprover intro: assms THE_defaultI') |
|
1724 |
|
1725 lemma the_default_eqvt: |
|
1726 assumes unique: "\<exists>!x. P x" |
|
1727 shows "(p \<bullet> (THE_default d P)) = (THE_default (p \<bullet> d) (p \<bullet> P))" |
|
1728 apply(rule THE_default1_equality [symmetric]) |
|
1729 apply(rule_tac p="-p" in permute_boolE) |
|
1730 apply(simp add: ex1_eqvt) |
|
1731 apply(rule unique) |
|
1732 apply(rule_tac p="-p" in permute_boolE) |
|
1733 apply(rule subst[OF permute_fun_app_eq]) |
|
1734 apply(simp) |
|
1735 apply(rule THE_defaultI'[OF unique]) |
|
1736 done |
|
1737 |
|
1738 lemma fundef_ex1_eqvt: |
|
1739 fixes x::"'a::pt" |
|
1740 assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (G x))" |
|
1741 assumes eqvt: "eqvt G" |
|
1742 assumes ex1: "\<exists>!y. G x y" |
|
1743 shows "(p \<bullet> (f x)) = f (p \<bullet> x)" |
|
1744 apply(simp only: f_def) |
|
1745 apply(subst the_default_eqvt) |
|
1746 apply(rule ex1) |
|
1747 apply(rule THE_default1_equality [symmetric]) |
|
1748 apply(rule_tac p="-p" in permute_boolE) |
|
1749 apply(perm_simp add: permute_minus_cancel) |
|
1750 using eqvt[simplified eqvt_def] |
|
1751 apply(simp) |
|
1752 apply(rule ex1) |
|
1753 apply(rule THE_defaultI2) |
|
1754 apply(rule_tac p="-p" in permute_boolE) |
|
1755 apply(perm_simp add: permute_minus_cancel) |
|
1756 apply(rule ex1) |
|
1757 apply(perm_simp) |
|
1758 using eqvt[simplified eqvt_def] |
|
1759 apply(simp) |
|
1760 done |
|
1761 |
|
1762 lemma fundef_ex1_eqvt_at: |
|
1763 fixes x::"'a::pt" |
|
1764 assumes f_def: "f == (\<lambda>x::'a. THE_default (d x) (G x))" |
|
1765 assumes eqvt: "eqvt G" |
|
1766 assumes ex1: "\<exists>!y. G x y" |
|
1767 shows "eqvt_at f x" |
|
1768 unfolding eqvt_at_def |
|
1769 using assms |
|
1770 by (auto intro: fundef_ex1_eqvt) |
|
1771 |
|
1772 lemma fundef_ex1_prop: |
|
1773 fixes x::"'a::pt" |
|
1774 assumes f_def: "f \<equiv> (\<lambda>x::'a. THE_default (d x) (G x))" |
|
1775 assumes P_all: "\<And>x y. G x y \<Longrightarrow> P x y" |
|
1776 assumes ex1: "\<exists>!y. G x y" |
|
1777 shows "P x (f x)" |
|
1778 unfolding f_def |
|
1779 using ex1 |
|
1780 apply(erule_tac ex1E) |
|
1781 apply(rule THE_defaultI2) |
|
1782 apply(blast) |
|
1783 apply(rule P_all) |
|
1784 apply(assumption) |
|
1785 done |
|
1786 |
|
1787 |
|
1788 section {* Support of Finite Sets of Finitely Supported Elements *} |
|
1789 |
|
1790 text {* support and freshness for atom sets *} |
|
1791 |
|
1792 lemma supp_finite_atom_set: |
|
1793 fixes S::"atom set" |
|
1794 assumes "finite S" |
|
1795 shows "supp S = S" |
|
1796 apply(rule finite_supp_unique) |
|
1797 apply(simp add: supports_def) |
|
1798 apply(simp add: swap_set_not_in) |
|
1799 apply(rule assms) |
|
1800 apply(simp add: swap_set_in) |
|
1801 done |
|
1802 |
|
1803 lemma supp_cofinite_atom_set: |
|
1804 fixes S::"atom set" |
|
1805 assumes "finite (UNIV - S)" |
|
1806 shows "supp S = (UNIV - S)" |
|
1807 apply(rule finite_supp_unique) |
|
1808 apply(simp add: supports_def) |
|
1809 apply(simp add: swap_set_both_in) |
|
1810 apply(rule assms) |
|
1811 apply(subst swap_commute) |
|
1812 apply(simp add: swap_set_in) |
|
1813 done |
|
1814 |
|
1815 lemma fresh_finite_atom_set: |
|
1816 fixes S::"atom set" |
|
1817 assumes "finite S" |
|
1818 shows "a \<sharp> S \<longleftrightarrow> a \<notin> S" |
|
1819 unfolding fresh_def |
|
1820 by (simp add: supp_finite_atom_set[OF assms]) |
|
1821 |
|
1822 lemma fresh_minus_atom_set: |
|
1823 fixes S::"atom set" |
|
1824 assumes "finite S" |
|
1825 shows "a \<sharp> S - T \<longleftrightarrow> (a \<notin> T \<longrightarrow> a \<sharp> S)" |
|
1826 unfolding fresh_def |
|
1827 by (auto simp add: supp_finite_atom_set assms) |
|
1828 |
|
1829 lemma Union_supports_set: |
|
1830 shows "(\<Union>x \<in> S. supp x) supports S" |
|
1831 proof - |
|
1832 { fix a b |
|
1833 have "\<forall>x \<in> S. (a \<rightleftharpoons> b) \<bullet> x = x \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> S = S" |
|
1834 unfolding permute_set_def by force |
|
1835 } |
|
1836 then show "(\<Union>x \<in> S. supp x) supports S" |
|
1837 unfolding supports_def |
|
1838 by (simp add: fresh_def[symmetric] swap_fresh_fresh) |
|
1839 qed |
|
1840 |
|
1841 lemma Union_of_finite_supp_sets: |
|
1842 fixes S::"('a::fs set)" |
|
1843 assumes fin: "finite S" |
|
1844 shows "finite (\<Union>x\<in>S. supp x)" |
|
1845 using fin by (induct) (auto simp add: finite_supp) |
|
1846 |
|
1847 lemma Union_included_in_supp: |
|
1848 fixes S::"('a::fs set)" |
|
1849 assumes fin: "finite S" |
|
1850 shows "(\<Union>x\<in>S. supp x) \<subseteq> supp S" |
|
1851 proof - |
|
1852 have eqvt: "eqvt (\<lambda>S. \<Union> supp ` S)" |
|
1853 unfolding eqvt_def |
|
1854 by (perm_simp) (simp) |
|
1855 have "(\<Union>x\<in>S. supp x) = supp (\<Union>x\<in>S. supp x)" |
|
1856 by (rule supp_finite_atom_set[symmetric]) (rule Union_of_finite_supp_sets[OF fin]) |
|
1857 also have "\<dots> = supp ((\<lambda>S. \<Union> supp ` S) S)" by simp |
|
1858 also have "\<dots> \<subseteq> supp S" using eqvt |
|
1859 by (rule supp_fun_app_eqvt) |
|
1860 finally show "(\<Union>x\<in>S. supp x) \<subseteq> supp S" . |
|
1861 qed |
|
1862 |
|
1863 lemma supp_of_finite_sets: |
|
1864 fixes S::"('a::fs set)" |
|
1865 assumes fin: "finite S" |
|
1866 shows "(supp S) = (\<Union>x\<in>S. supp x)" |
|
1867 apply(rule subset_antisym) |
|
1868 apply(rule supp_is_subset) |
|
1869 apply(rule Union_supports_set) |
|
1870 apply(rule Union_of_finite_supp_sets[OF fin]) |
|
1871 apply(rule Union_included_in_supp[OF fin]) |
|
1872 done |
|
1873 |
|
1874 lemma finite_sets_supp: |
|
1875 fixes S::"('a::fs set)" |
|
1876 assumes "finite S" |
|
1877 shows "finite (supp S)" |
|
1878 using assms |
|
1879 by (simp only: supp_of_finite_sets Union_of_finite_supp_sets) |
|
1880 |
|
1881 lemma supp_of_finite_union: |
|
1882 fixes S T::"('a::fs) set" |
|
1883 assumes fin1: "finite S" |
|
1884 and fin2: "finite T" |
|
1885 shows "supp (S \<union> T) = supp S \<union> supp T" |
|
1886 using fin1 fin2 |
|
1887 by (simp add: supp_of_finite_sets) |
|
1888 |
|
1889 lemma supp_of_finite_insert: |
|
1890 fixes S::"('a::fs) set" |
|
1891 assumes fin: "finite S" |
|
1892 shows "supp (insert x S) = supp x \<union> supp S" |
|
1893 using fin |
|
1894 by (simp add: supp_of_finite_sets) |
|
1895 |
|
1896 lemma fresh_finite_insert: |
|
1897 fixes S::"('a::fs) set" |
|
1898 assumes fin: "finite S" |
|
1899 shows "a \<sharp> (insert x S) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S" |
|
1900 using fin unfolding fresh_def |
|
1901 by (simp add: supp_of_finite_insert) |
|
1902 |
|
1903 lemma supp_set_empty: |
|
1904 shows "supp {} = {}" |
|
1905 unfolding supp_def |
|
1906 by (simp add: empty_eqvt) |
|
1907 |
|
1908 lemma fresh_set_empty: |
|
1909 shows "a \<sharp> {}" |
|
1910 by (simp add: fresh_def supp_set_empty) |
|
1911 |
|
1912 lemma supp_set: |
|
1913 fixes xs :: "('a::fs) list" |
|
1914 shows "supp (set xs) = supp xs" |
|
1915 apply(induct xs) |
|
1916 apply(simp add: supp_set_empty supp_Nil) |
|
1917 apply(simp add: supp_Cons supp_of_finite_insert) |
|
1918 done |
|
1919 |
|
1920 lemma fresh_set: |
|
1921 fixes xs :: "('a::fs) list" |
|
1922 shows "a \<sharp> (set xs) \<longleftrightarrow> a \<sharp> xs" |
|
1923 unfolding fresh_def |
|
1924 by (simp add: supp_set) |
|
1925 |
|
1926 |
|
1927 subsection {* Type @{typ "'a multiset"} is finitely supported *} |
|
1928 |
|
1929 lemma set_of_eqvt[eqvt]: |
|
1930 shows "p \<bullet> (set_of M) = set_of (p \<bullet> M)" |
|
1931 by (induct M) (simp_all add: insert_eqvt empty_eqvt) |
|
1932 |
|
1933 lemma supp_set_of: |
|
1934 shows "supp (set_of M) \<subseteq> supp M" |
|
1935 apply (rule supp_fun_app_eqvt) |
|
1936 unfolding eqvt_def |
|
1937 apply(perm_simp) |
|
1938 apply(simp) |
|
1939 done |
|
1940 |
|
1941 lemma Union_finite_multiset: |
|
1942 fixes M::"'a::fs multiset" |
|
1943 shows "finite (\<Union>{supp x | x. x \<in># M})" |
|
1944 proof - |
|
1945 have "finite (\<Union>(supp ` {x. x \<in># M}))" |
|
1946 by (induct M) (simp_all add: Collect_imp_eq Collect_neg_eq finite_supp) |
|
1947 then show "finite (\<Union>{supp x | x. x \<in># M})" |
|
1948 by (simp only: image_Collect) |
|
1949 qed |
|
1950 |
|
1951 lemma Union_supports_multiset: |
|
1952 shows "\<Union>{supp x | x. x :# M} supports M" |
|
1953 proof - |
|
1954 have sw: "\<And>a b. ((\<And>x. x :# M \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x) \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> M = M)" |
|
1955 unfolding permute_multiset_def |
|
1956 apply(induct M) |
|
1957 apply(simp_all) |
|
1958 done |
|
1959 show "(\<Union>{supp x | x. x :# M}) supports M" |
|
1960 unfolding supports_def |
|
1961 apply(clarify) |
|
1962 apply(rule sw) |
|
1963 apply(rule swap_fresh_fresh) |
|
1964 apply(simp_all only: fresh_def) |
|
1965 apply(auto) |
|
1966 apply(metis neq0_conv)+ |
|
1967 done |
|
1968 qed |
|
1969 |
|
1970 lemma Union_included_multiset: |
|
1971 fixes M::"('a::fs multiset)" |
|
1972 shows "(\<Union>{supp x | x. x \<in># M}) \<subseteq> supp M" |
|
1973 proof - |
|
1974 have "(\<Union>{supp x | x. x \<in># M}) = (\<Union>{supp x | x. x \<in> set_of M})" by simp |
|
1975 also have "... \<subseteq> (\<Union>x \<in> set_of M. supp x)" by auto |
|
1976 also have "... = supp (set_of M)" by (simp add: subst supp_of_finite_sets) |
|
1977 also have " ... \<subseteq> supp M" by (rule supp_set_of) |
|
1978 finally show "(\<Union>{supp x | x. x \<in># M}) \<subseteq> supp M" . |
|
1979 qed |
|
1980 |
|
1981 lemma supp_of_multisets: |
|
1982 fixes M::"('a::fs multiset)" |
|
1983 shows "(supp M) = (\<Union>{supp x | x. x :# M})" |
|
1984 apply(rule subset_antisym) |
|
1985 apply(rule supp_is_subset) |
|
1986 apply(rule Union_supports_multiset) |
|
1987 apply(rule Union_finite_multiset) |
|
1988 apply(rule Union_included_multiset) |
|
1989 done |
|
1990 |
|
1991 lemma multisets_supp_finite: |
|
1992 fixes M::"('a::fs multiset)" |
|
1993 shows "finite (supp M)" |
|
1994 by (simp only: supp_of_multisets Union_finite_multiset) |
|
1995 |
|
1996 lemma supp_of_multiset_union: |
|
1997 fixes M N::"('a::fs) multiset" |
|
1998 shows "supp (M + N) = supp M \<union> supp N" |
|
1999 by (auto simp add: supp_of_multisets) |
|
2000 |
|
2001 lemma supp_empty_mset [simp]: |
|
2002 shows "supp {#} = {}" |
|
2003 unfolding supp_def |
|
2004 by simp |
|
2005 |
|
2006 instance multiset :: (fs) fs |
|
2007 apply (default) |
|
2008 apply (rule multisets_supp_finite) |
|
2009 done |
|
2010 |
|
2011 subsection {* Type @{typ "'a fset"} is finitely supported *} |
|
2012 |
|
2013 lemma supp_fset [simp]: |
|
2014 shows "supp (fset S) = supp S" |
|
2015 unfolding supp_def |
|
2016 by (simp add: fset_eqvt fset_cong) |
|
2017 |
|
2018 lemma supp_empty_fset [simp]: |
|
2019 shows "supp {||} = {}" |
|
2020 unfolding supp_def |
|
2021 by simp |
|
2022 |
|
2023 lemma fresh_empty_fset: |
|
2024 shows "a \<sharp> {||}" |
|
2025 unfolding fresh_def |
|
2026 by (simp) |
|
2027 |
|
2028 lemma supp_insert_fset [simp]: |
|
2029 fixes x::"'a::fs" |
|
2030 and S::"'a fset" |
|
2031 shows "supp (insert_fset x S) = supp x \<union> supp S" |
|
2032 apply(subst supp_fset[symmetric]) |
|
2033 apply(simp add: supp_of_finite_insert) |
|
2034 done |
|
2035 |
|
2036 lemma fresh_insert_fset: |
|
2037 fixes x::"'a::fs" |
|
2038 and S::"'a fset" |
|
2039 shows "a \<sharp> insert_fset x S \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> S" |
|
2040 unfolding fresh_def |
|
2041 by (simp) |
|
2042 |
|
2043 lemma fset_finite_supp: |
|
2044 fixes S::"('a::fs) fset" |
|
2045 shows "finite (supp S)" |
|
2046 by (induct S) (simp_all add: finite_supp) |
|
2047 |
|
2048 lemma supp_union_fset: |
|
2049 fixes S T::"'a::fs fset" |
|
2050 shows "supp (S |\<union>| T) = supp S \<union> supp T" |
|
2051 by (induct S) (auto) |
|
2052 |
|
2053 lemma fresh_union_fset: |
|
2054 fixes S T::"'a::fs fset" |
|
2055 shows "a \<sharp> S |\<union>| T \<longleftrightarrow> a \<sharp> S \<and> a \<sharp> T" |
|
2056 unfolding fresh_def |
|
2057 by (simp add: supp_union_fset) |
|
2058 |
|
2059 instance fset :: (fs) fs |
|
2060 apply (default) |
|
2061 apply (rule fset_finite_supp) |
|
2062 done |
|
2063 |
|
2064 |
|
2065 section {* Freshness and Fresh-Star *} |
|
2066 |
|
2067 lemma fresh_Unit_elim: |
|
2068 shows "(a \<sharp> () \<Longrightarrow> PROP C) \<equiv> PROP C" |
|
2069 by (simp add: fresh_Unit) |
|
2070 |
|
2071 lemma fresh_Pair_elim: |
|
2072 shows "(a \<sharp> (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> a \<sharp> y \<Longrightarrow> PROP C)" |
|
2073 by rule (simp_all add: fresh_Pair) |
|
2074 |
|
2075 (* this rule needs to be added before the fresh_prodD is *) |
|
2076 (* added to the simplifier with mksimps *) |
|
2077 lemma [simp]: |
|
2078 shows "a \<sharp> x1 \<Longrightarrow> a \<sharp> x2 \<Longrightarrow> a \<sharp> (x1, x2)" |
|
2079 by (simp add: fresh_Pair) |
|
2080 |
|
2081 lemma fresh_PairD: |
|
2082 shows "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> x" |
|
2083 and "a \<sharp> (x, y) \<Longrightarrow> a \<sharp> y" |
|
2084 by (simp_all add: fresh_Pair) |
|
2085 |
|
2086 declaration {* fn _ => |
|
2087 let |
|
2088 val mksimps_pairs = (@{const_name Nominal2_Base.fresh}, @{thms fresh_PairD}) :: mksimps_pairs |
|
2089 in |
|
2090 Simplifier.map_ss (fn ss => Simplifier.set_mksimps (mksimps mksimps_pairs) ss) |
|
2091 end |
|
2092 *} |
|
2093 |
|
2094 text {* The fresh-star generalisation of fresh is used in strong |
|
2095 induction principles. *} |
|
2096 |
|
2097 definition |
|
2098 fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80) |
|
2099 where |
|
2100 "as \<sharp>* x \<equiv> \<forall>a \<in> as. a \<sharp> x" |
|
2101 |
|
2102 lemma fresh_star_supp_conv: |
|
2103 shows "supp x \<sharp>* y \<Longrightarrow> supp y \<sharp>* x" |
|
2104 by (auto simp add: fresh_star_def fresh_def) |
|
2105 |
|
2106 lemma fresh_star_perm_set_conv: |
|
2107 fixes p::"perm" |
|
2108 assumes fresh: "as \<sharp>* p" |
|
2109 and fin: "finite as" |
|
2110 shows "supp p \<sharp>* as" |
|
2111 apply(rule fresh_star_supp_conv) |
|
2112 apply(simp add: supp_finite_atom_set fin fresh) |
|
2113 done |
|
2114 |
|
2115 lemma fresh_star_atom_set_conv: |
|
2116 assumes fresh: "as \<sharp>* bs" |
|
2117 and fin: "finite as" "finite bs" |
|
2118 shows "bs \<sharp>* as" |
|
2119 using fresh |
|
2120 unfolding fresh_star_def fresh_def |
|
2121 by (auto simp add: supp_finite_atom_set fin) |
|
2122 |
|
2123 lemma atom_fresh_star_disjoint: |
|
2124 assumes fin: "finite bs" |
|
2125 shows "as \<sharp>* bs \<longleftrightarrow> (as \<inter> bs = {})" |
|
2126 |
|
2127 unfolding fresh_star_def fresh_def |
|
2128 by (auto simp add: supp_finite_atom_set fin) |
|
2129 |
|
2130 |
|
2131 lemma fresh_star_Pair: |
|
2132 shows "as \<sharp>* (x, y) = (as \<sharp>* x \<and> as \<sharp>* y)" |
|
2133 by (auto simp add: fresh_star_def fresh_Pair) |
|
2134 |
|
2135 lemma fresh_star_list: |
|
2136 shows "as \<sharp>* (xs @ ys) \<longleftrightarrow> as \<sharp>* xs \<and> as \<sharp>* ys" |
|
2137 and "as \<sharp>* (x # xs) \<longleftrightarrow> as \<sharp>* x \<and> as \<sharp>* xs" |
|
2138 and "as \<sharp>* []" |
|
2139 by (auto simp add: fresh_star_def fresh_Nil fresh_Cons fresh_append) |
|
2140 |
|
2141 lemma fresh_star_set: |
|
2142 fixes xs::"('a::fs) list" |
|
2143 shows "as \<sharp>* set xs \<longleftrightarrow> as \<sharp>* xs" |
|
2144 unfolding fresh_star_def |
|
2145 by (simp add: fresh_set) |
|
2146 |
|
2147 lemma fresh_star_singleton: |
|
2148 fixes a::"atom" |
|
2149 shows "as \<sharp>* {a} \<longleftrightarrow> as \<sharp>* a" |
|
2150 by (simp add: fresh_star_def fresh_finite_insert fresh_set_empty) |
|
2151 |
|
2152 lemma fresh_star_fset: |
|
2153 fixes xs::"('a::fs) list" |
|
2154 shows "as \<sharp>* fset S \<longleftrightarrow> as \<sharp>* S" |
|
2155 by (simp add: fresh_star_def fresh_def) |
|
2156 |
|
2157 lemma fresh_star_Un: |
|
2158 shows "(as \<union> bs) \<sharp>* x = (as \<sharp>* x \<and> bs \<sharp>* x)" |
|
2159 by (auto simp add: fresh_star_def) |
|
2160 |
|
2161 lemma fresh_star_insert: |
|
2162 shows "(insert a as) \<sharp>* x = (a \<sharp> x \<and> as \<sharp>* x)" |
|
2163 by (auto simp add: fresh_star_def) |
|
2164 |
|
2165 lemma fresh_star_Un_elim: |
|
2166 "((as \<union> bs) \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (as \<sharp>* x \<Longrightarrow> bs \<sharp>* x \<Longrightarrow> PROP C)" |
|
2167 unfolding fresh_star_def |
|
2168 apply(rule) |
|
2169 apply(erule meta_mp) |
|
2170 apply(auto) |
|
2171 done |
|
2172 |
|
2173 lemma fresh_star_insert_elim: |
|
2174 "(insert a as \<sharp>* x \<Longrightarrow> PROP C) \<equiv> (a \<sharp> x \<Longrightarrow> as \<sharp>* x \<Longrightarrow> PROP C)" |
|
2175 unfolding fresh_star_def |
|
2176 by rule (simp_all add: fresh_star_def) |
|
2177 |
|
2178 lemma fresh_star_empty_elim: |
|
2179 "({} \<sharp>* x \<Longrightarrow> PROP C) \<equiv> PROP C" |
|
2180 by (simp add: fresh_star_def) |
|
2181 |
|
2182 lemma fresh_star_Unit_elim: |
|
2183 shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C" |
|
2184 by (simp add: fresh_star_def fresh_Unit) |
|
2185 |
|
2186 lemma fresh_star_Pair_elim: |
|
2187 shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)" |
|
2188 by (rule, simp_all add: fresh_star_Pair) |
|
2189 |
|
2190 lemma fresh_star_zero: |
|
2191 shows "as \<sharp>* (0::perm)" |
|
2192 unfolding fresh_star_def |
|
2193 by (simp add: fresh_zero_perm) |
|
2194 |
|
2195 lemma fresh_star_plus: |
|
2196 fixes p q::perm |
|
2197 shows "\<lbrakk>a \<sharp>* p; a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)" |
|
2198 unfolding fresh_star_def |
|
2199 by (simp add: fresh_plus_perm) |
|
2200 |
|
2201 lemma fresh_star_permute_iff: |
|
2202 shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x" |
|
2203 unfolding fresh_star_def |
|
2204 by (metis mem_permute_iff permute_minus_cancel(1) fresh_permute_iff) |
|
2205 |
|
2206 lemma fresh_star_eqvt [eqvt]: |
|
2207 shows "p \<bullet> (as \<sharp>* x) \<longleftrightarrow> (p \<bullet> as) \<sharp>* (p \<bullet> x)" |
|
2208 unfolding fresh_star_def |
|
2209 by (perm_simp) (rule refl) |
|
2210 |
|
2211 |
|
2212 |
|
2213 section {* Induction principle for permutations *} |
|
2214 |
|
2215 lemma smaller_supp: |
|
2216 assumes a: "a \<in> supp p" |
|
2217 shows "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p" |
|
2218 proof - |
|
2219 have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subseteq> supp p" |
|
2220 unfolding supp_perm by (auto simp add: swap_atom) |
|
2221 moreover |
|
2222 have "a \<notin> supp ((p \<bullet> a \<rightleftharpoons> a) + p)" by (simp add: supp_perm) |
|
2223 then have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<noteq> supp p" using a by auto |
|
2224 ultimately |
|
2225 show "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p" by auto |
|
2226 qed |
|
2227 |
|
2228 |
|
2229 lemma perm_struct_induct[consumes 1, case_names zero swap]: |
|
2230 assumes S: "supp p \<subseteq> S" |
|
2231 and zero: "P 0" |
|
2232 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)" |
|
2233 shows "P p" |
|
2234 proof - |
|
2235 have "finite (supp p)" by (simp add: finite_supp) |
|
2236 then show "P p" using S |
|
2237 proof(induct A\<equiv>"supp p" arbitrary: p rule: finite_psubset_induct) |
|
2238 case (psubset p) |
|
2239 then have ih: "\<And>q. supp q \<subset> supp p \<Longrightarrow> P q" by auto |
|
2240 have as: "supp p \<subseteq> S" by fact |
|
2241 { assume "supp p = {}" |
|
2242 then have "p = 0" by (simp add: supp_perm perm_eq_iff) |
|
2243 then have "P p" using zero by simp |
|
2244 } |
|
2245 moreover |
|
2246 { assume "supp p \<noteq> {}" |
|
2247 then obtain a where a0: "a \<in> supp p" by blast |
|
2248 then have a1: "p \<bullet> a \<in> S" "a \<in> S" "sort_of (p \<bullet> a) = sort_of a" "p \<bullet> a \<noteq> a" |
|
2249 using as by (auto simp add: supp_atom supp_perm swap_atom) |
|
2250 let ?q = "(p \<bullet> a \<rightleftharpoons> a) + p" |
|
2251 have a2: "supp ?q \<subset> supp p" using a0 smaller_supp by simp |
|
2252 then have "P ?q" using ih by simp |
|
2253 moreover |
|
2254 have "supp ?q \<subseteq> S" using as a2 by simp |
|
2255 ultimately have "P ((p \<bullet> a \<rightleftharpoons> a) + ?q)" using as a1 swap by simp |
|
2256 moreover |
|
2257 have "p = (p \<bullet> a \<rightleftharpoons> a) + ?q" by (simp add: perm_eq_iff) |
|
2258 ultimately have "P p" by simp |
|
2259 } |
|
2260 ultimately show "P p" by blast |
|
2261 qed |
|
2262 qed |
|
2263 |
|
2264 lemma perm_simple_struct_induct[case_names zero swap]: |
|
2265 assumes zero: "P 0" |
|
2266 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)" |
|
2267 shows "P p" |
|
2268 by (rule_tac S="supp p" in perm_struct_induct) |
|
2269 (auto intro: zero swap) |
|
2270 |
|
2271 lemma perm_struct_induct2[consumes 1, case_names zero swap plus]: |
|
2272 assumes S: "supp p \<subseteq> S" |
|
2273 assumes zero: "P 0" |
|
2274 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)" |
|
2275 assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2; supp p1 \<subseteq> S; supp p2 \<subseteq> S\<rbrakk> \<Longrightarrow> P (p1 + p2)" |
|
2276 shows "P p" |
|
2277 using S |
|
2278 by (induct p rule: perm_struct_induct) |
|
2279 (auto intro: zero plus swap simp add: supp_swap) |
|
2280 |
|
2281 lemma perm_simple_struct_induct2[case_names zero swap plus]: |
|
2282 assumes zero: "P 0" |
|
2283 assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)" |
|
2284 assumes plus: "\<And>p1 p2. \<lbrakk>P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)" |
|
2285 shows "P p" |
|
2286 by (rule_tac S="supp p" in perm_struct_induct2) |
|
2287 (auto intro: zero swap plus) |
|
2288 |
|
2289 lemma supp_perm_singleton: |
|
2290 fixes p::"perm" |
|
2291 shows "supp p \<subseteq> {b} \<longleftrightarrow> p = 0" |
|
2292 proof - |
|
2293 { assume "supp p \<subseteq> {b}" |
|
2294 then have "p = 0" |
|
2295 by (induct p rule: perm_struct_induct) (simp_all) |
|
2296 } |
|
2297 then show "supp p \<subseteq> {b} \<longleftrightarrow> p = 0" by (auto simp add: supp_zero_perm) |
|
2298 qed |
|
2299 |
|
2300 lemma supp_perm_pair: |
|
2301 fixes p::"perm" |
|
2302 shows "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)" |
|
2303 proof - |
|
2304 { assume "supp p \<subseteq> {a, b}" |
|
2305 then have "p = 0 \<or> p = (b \<rightleftharpoons> a)" |
|
2306 apply (induct p rule: perm_struct_induct) |
|
2307 apply (auto simp add: swap_cancel supp_zero_perm supp_swap) |
|
2308 apply (simp add: swap_commute) |
|
2309 done |
|
2310 } |
|
2311 then show "supp p \<subseteq> {a, b} \<longleftrightarrow> p = 0 \<or> p = (b \<rightleftharpoons> a)" |
|
2312 by (auto simp add: supp_zero_perm supp_swap split: if_splits) |
|
2313 qed |
|
2314 |
|
2315 lemma supp_perm_eq: |
|
2316 assumes "(supp x) \<sharp>* p" |
|
2317 shows "p \<bullet> x = x" |
|
2318 proof - |
|
2319 from assms have "supp p \<subseteq> {a. a \<sharp> x}" |
|
2320 unfolding supp_perm fresh_star_def fresh_def by auto |
|
2321 then show "p \<bullet> x = x" |
|
2322 proof (induct p rule: perm_struct_induct) |
|
2323 case zero |
|
2324 show "0 \<bullet> x = x" by simp |
|
2325 next |
|
2326 case (swap p a b) |
|
2327 then have "a \<sharp> x" "b \<sharp> x" "p \<bullet> x = x" by simp_all |
|
2328 then show "((a \<rightleftharpoons> b) + p) \<bullet> x = x" by (simp add: swap_fresh_fresh) |
|
2329 qed |
|
2330 qed |
|
2331 |
|
2332 text {* same lemma as above, but proved with a different induction principle *} |
|
2333 lemma supp_perm_eq_test: |
|
2334 assumes "(supp x) \<sharp>* p" |
|
2335 shows "p \<bullet> x = x" |
|
2336 proof - |
|
2337 from assms have "supp p \<subseteq> {a. a \<sharp> x}" |
|
2338 unfolding supp_perm fresh_star_def fresh_def by auto |
|
2339 then show "p \<bullet> x = x" |
|
2340 proof (induct p rule: perm_struct_induct2) |
|
2341 case zero |
|
2342 show "0 \<bullet> x = x" by simp |
|
2343 next |
|
2344 case (swap a b) |
|
2345 then have "a \<sharp> x" "b \<sharp> x" by simp_all |
|
2346 then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh) |
|
2347 next |
|
2348 case (plus p1 p2) |
|
2349 have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+ |
|
2350 then show "(p1 + p2) \<bullet> x = x" by simp |
|
2351 qed |
|
2352 qed |
|
2353 |
|
2354 lemma perm_supp_eq: |
|
2355 assumes a: "(supp p) \<sharp>* x" |
|
2356 shows "p \<bullet> x = x" |
|
2357 proof - |
|
2358 from assms have "supp p \<subseteq> {a. a \<sharp> x}" |
|
2359 unfolding supp_perm fresh_star_def fresh_def by auto |
|
2360 then show "p \<bullet> x = x" |
|
2361 proof (induct p rule: perm_struct_induct2) |
|
2362 case zero |
|
2363 show "0 \<bullet> x = x" by simp |
|
2364 next |
|
2365 case (swap a b) |
|
2366 then have "a \<sharp> x" "b \<sharp> x" by simp_all |
|
2367 then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh) |
|
2368 next |
|
2369 case (plus p1 p2) |
|
2370 have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+ |
|
2371 then show "(p1 + p2) \<bullet> x = x" by simp |
|
2372 qed |
|
2373 qed |
|
2374 |
|
2375 |
|
2376 |
|
2377 |
|
2378 lemma supp_perm_perm_eq: |
|
2379 assumes a: "\<forall>a \<in> supp x. p \<bullet> a = q \<bullet> a" |
|
2380 shows "p \<bullet> x = q \<bullet> x" |
|
2381 proof - |
|
2382 from a have "\<forall>a \<in> supp x. (-q + p) \<bullet> a = a" by simp |
|
2383 then have "\<forall>a \<in> supp x. a \<notin> supp (-q + p)" |
|
2384 unfolding supp_perm by simp |
|
2385 then have "supp x \<sharp>* (-q + p)" |
|
2386 unfolding fresh_star_def fresh_def by simp |
|
2387 then have "(-q + p) \<bullet> x = x" by (simp only: supp_perm_eq) |
|
2388 then show "p \<bullet> x = q \<bullet> x" |
|
2389 by (metis permute_minus_cancel(1) permute_plus) |
|
2390 qed |
|
2391 |
|
2392 text {* disagreement set *} |
|
2393 |
|
2394 definition |
|
2395 dset :: "perm \<Rightarrow> perm \<Rightarrow> atom set" |
|
2396 where |
|
2397 "dset p q = {a::atom. p \<bullet> a \<noteq> q \<bullet> a}" |
|
2398 |
|
2399 lemma ds_fresh: |
|
2400 assumes "dset p q \<sharp>* x" |
|
2401 shows "p \<bullet> x = q \<bullet> x" |
|
2402 using assms |
|
2403 unfolding dset_def fresh_star_def fresh_def |
|
2404 by (auto intro: supp_perm_perm_eq) |
|
2405 |
|
2406 lemma atom_set_perm_eq: |
|
2407 assumes a: "as \<sharp>* p" |
|
2408 shows "p \<bullet> as = as" |
|
2409 proof - |
|
2410 from a have "supp p \<subseteq> {a. a \<notin> as}" |
|
2411 unfolding supp_perm fresh_star_def fresh_def by auto |
|
2412 then show "p \<bullet> as = as" |
|
2413 proof (induct p rule: perm_struct_induct) |
|
2414 case zero |
|
2415 show "0 \<bullet> as = as" by simp |
|
2416 next |
|
2417 case (swap p a b) |
|
2418 then have "a \<notin> as" "b \<notin> as" "p \<bullet> as = as" by simp_all |
|
2419 then show "((a \<rightleftharpoons> b) + p) \<bullet> as = as" by (simp add: swap_set_not_in) |
|
2420 qed |
|
2421 qed |
|
2422 |
|
2423 section {* Avoiding of atom sets *} |
|
2424 |
|
2425 text {* |
|
2426 For every set of atoms, there is another set of atoms |
|
2427 avoiding a finitely supported c and there is a permutation |
|
2428 which 'translates' between both sets. |
|
2429 *} |
|
2430 |
|
2431 lemma at_set_avoiding_aux: |
|
2432 fixes Xs::"atom set" |
|
2433 and As::"atom set" |
|
2434 assumes b: "Xs \<subseteq> As" |
|
2435 and c: "finite As" |
|
2436 shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) = (Xs \<union> (p \<bullet> Xs))" |
|
2437 proof - |
|
2438 from b c have "finite Xs" by (rule finite_subset) |
|
2439 then show ?thesis using b |
|
2440 proof (induct rule: finite_subset_induct) |
|
2441 case empty |
|
2442 have "0 \<bullet> {} \<inter> As = {}" by simp |
|
2443 moreover |
|
2444 have "supp (0::perm) = {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm) |
|
2445 ultimately show ?case by blast |
|
2446 next |
|
2447 case (insert x Xs) |
|
2448 then obtain p where |
|
2449 p1: "(p \<bullet> Xs) \<inter> As = {}" and |
|
2450 p2: "supp p = (Xs \<union> (p \<bullet> Xs))" by blast |
|
2451 from `x \<in> As` p1 have "x \<notin> p \<bullet> Xs" by fast |
|
2452 with `x \<notin> Xs` p2 have "x \<notin> supp p" by fast |
|
2453 hence px: "p \<bullet> x = x" unfolding supp_perm by simp |
|
2454 have "finite (As \<union> p \<bullet> Xs \<union> supp p)" |
|
2455 using `finite As` `finite Xs` |
|
2456 by (simp add: permute_set_eq_image finite_supp) |
|
2457 then obtain y where "y \<notin> (As \<union> p \<bullet> Xs \<union> supp p)" "sort_of y = sort_of x" |
|
2458 by (rule obtain_atom) |
|
2459 hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "y \<notin> supp p" "sort_of y = sort_of x" |
|
2460 by simp_all |
|
2461 hence py: "p \<bullet> y = y" "x \<noteq> y" using `x \<in> As` |
|
2462 by (auto simp add: supp_perm) |
|
2463 let ?q = "(x \<rightleftharpoons> y) + p" |
|
2464 have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)" |
|
2465 unfolding insert_eqvt |
|
2466 using `p \<bullet> x = x` `sort_of y = sort_of x` |
|
2467 using `x \<notin> p \<bullet> Xs` `y \<notin> p \<bullet> Xs` |
|
2468 by (simp add: swap_atom swap_set_not_in) |
|
2469 have "?q \<bullet> insert x Xs \<inter> As = {}" |
|
2470 using `y \<notin> As` `p \<bullet> Xs \<inter> As = {}` |
|
2471 unfolding q by simp |
|
2472 moreover |
|
2473 have "supp (x \<rightleftharpoons> y) \<inter> supp p = {}" using px py `sort_of y = sort_of x` |
|
2474 unfolding supp_swap by (simp add: supp_perm) |
|
2475 then have "supp ?q = (supp (x \<rightleftharpoons> y) \<union> supp p)" |
|
2476 by (simp add: supp_plus_perm_eq) |
|
2477 then have "supp ?q = insert x Xs \<union> ?q \<bullet> insert x Xs" |
|
2478 using p2 `sort_of y = sort_of x` `x \<noteq> y` unfolding q supp_swap |
|
2479 by auto |
|
2480 ultimately show ?case by blast |
|
2481 qed |
|
2482 qed |
|
2483 |
|
2484 lemma at_set_avoiding: |
|
2485 assumes a: "finite Xs" |
|
2486 and b: "finite (supp c)" |
|
2487 obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) = (Xs \<union> (p \<bullet> Xs))" |
|
2488 using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"] |
|
2489 unfolding fresh_star_def fresh_def by blast |
|
2490 |
|
2491 lemma at_set_avoiding1: |
|
2492 assumes "finite xs" |
|
2493 and "finite (supp c)" |
|
2494 shows "\<exists>p. (p \<bullet> xs) \<sharp>* c" |
|
2495 using assms |
|
2496 apply(erule_tac c="c" in at_set_avoiding) |
|
2497 apply(auto) |
|
2498 done |
|
2499 |
|
2500 lemma at_set_avoiding2: |
|
2501 assumes "finite xs" |
|
2502 and "finite (supp c)" "finite (supp x)" |
|
2503 and "xs \<sharp>* x" |
|
2504 shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p" |
|
2505 using assms |
|
2506 apply(erule_tac c="(c, x)" in at_set_avoiding) |
|
2507 apply(simp add: supp_Pair) |
|
2508 apply(rule_tac x="p" in exI) |
|
2509 apply(simp add: fresh_star_Pair) |
|
2510 apply(rule fresh_star_supp_conv) |
|
2511 apply(auto simp add: fresh_star_def) |
|
2512 done |
|
2513 |
|
2514 lemma at_set_avoiding3: |
|
2515 assumes "finite xs" |
|
2516 and "finite (supp c)" "finite (supp x)" |
|
2517 and "xs \<sharp>* x" |
|
2518 shows "\<exists>p. (p \<bullet> xs) \<sharp>* c \<and> supp x \<sharp>* p \<and> supp p = xs \<union> (p \<bullet> xs)" |
|
2519 using assms |
|
2520 apply(erule_tac c="(c, x)" in at_set_avoiding) |
|
2521 apply(simp add: supp_Pair) |
|
2522 apply(rule_tac x="p" in exI) |
|
2523 apply(simp add: fresh_star_Pair) |
|
2524 apply(rule fresh_star_supp_conv) |
|
2525 apply(auto simp add: fresh_star_def) |
|
2526 done |
|
2527 |
|
2528 lemma at_set_avoiding2_atom: |
|
2529 assumes "finite (supp c)" "finite (supp x)" |
|
2530 and b: "a \<sharp> x" |
|
2531 shows "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" |
|
2532 proof - |
|
2533 have a: "{a} \<sharp>* x" unfolding fresh_star_def by (simp add: b) |
|
2534 obtain p where p1: "(p \<bullet> {a}) \<sharp>* c" and p2: "supp x \<sharp>* p" |
|
2535 using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast |
|
2536 have c: "(p \<bullet> a) \<sharp> c" using p1 |
|
2537 unfolding fresh_star_def Ball_def |
|
2538 by(erule_tac x="p \<bullet> a" in allE) (simp add: permute_set_def) |
|
2539 hence "p \<bullet> a \<sharp> c \<and> supp x \<sharp>* p" using p2 by blast |
|
2540 then show "\<exists>p. (p \<bullet> a) \<sharp> c \<and> supp x \<sharp>* p" by blast |
|
2541 qed |
|
2542 |
|
2543 |
|
2544 section {* Renaming permutations *} |
|
2545 |
|
2546 lemma set_renaming_perm: |
|
2547 assumes b: "finite bs" |
|
2548 shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" |
|
2549 using b |
|
2550 proof (induct) |
|
2551 case empty |
|
2552 have "(\<forall>b \<in> {}. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> {} \<union> p \<bullet> {}" |
|
2553 by (simp add: permute_set_def supp_perm) |
|
2554 then show "\<exists>q. (\<forall>b \<in> {}. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> {} \<union> p \<bullet> {}" by blast |
|
2555 next |
|
2556 case (insert a bs) |
|
2557 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 |
|
2558 then obtain q where *: "\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b" and **: "supp q \<subseteq> bs \<union> p \<bullet> bs" |
|
2559 by auto |
|
2560 { assume 1: "q \<bullet> a = p \<bullet> a" |
|
2561 have "\<forall>b \<in> (insert a bs). q \<bullet> b = p \<bullet> b" using 1 * by simp |
|
2562 moreover |
|
2563 have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" |
|
2564 using ** by (auto simp add: insert_eqvt) |
|
2565 ultimately |
|
2566 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 |
|
2567 } |
|
2568 moreover |
|
2569 { assume 2: "q \<bullet> a \<noteq> p \<bullet> a" |
|
2570 def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q" |
|
2571 have "\<forall>b \<in> insert a bs. q' \<bullet> b = p \<bullet> b" using 2 * `a \<notin> bs` unfolding q'_def |
|
2572 by (auto simp add: swap_atom) |
|
2573 moreover |
|
2574 { have "{q \<bullet> a, p \<bullet> a} \<subseteq> insert a bs \<union> p \<bullet> insert a bs" |
|
2575 using ** |
|
2576 apply (auto simp add: supp_perm insert_eqvt) |
|
2577 apply (subgoal_tac "q \<bullet> a \<in> bs \<union> p \<bullet> bs") |
|
2578 apply(auto)[1] |
|
2579 apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}") |
|
2580 apply(blast) |
|
2581 apply(simp) |
|
2582 done |
|
2583 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) |
|
2584 moreover |
|
2585 have "supp q \<subseteq> insert a bs \<union> p \<bullet> insert a bs" |
|
2586 using ** by (auto simp add: insert_eqvt) |
|
2587 ultimately |
|
2588 have "supp q' \<subseteq> insert a bs \<union> p \<bullet> insert a bs" |
|
2589 unfolding q'_def using supp_plus_perm by blast |
|
2590 } |
|
2591 ultimately |
|
2592 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 |
|
2593 } |
|
2594 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" |
|
2595 by blast |
|
2596 qed |
|
2597 |
|
2598 lemma set_renaming_perm2: |
|
2599 shows "\<exists>q. (\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" |
|
2600 proof - |
|
2601 have "finite (bs \<inter> supp p)" by (simp add: finite_supp) |
|
2602 then obtain q |
|
2603 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))" |
|
2604 using set_renaming_perm by blast |
|
2605 from ** have "supp q \<subseteq> bs \<union> (p \<bullet> bs)" by (auto simp add: inter_eqvt) |
|
2606 moreover |
|
2607 have "\<forall>b \<in> bs - supp p. q \<bullet> b = p \<bullet> b" |
|
2608 apply(auto) |
|
2609 apply(subgoal_tac "b \<notin> supp q") |
|
2610 apply(simp add: fresh_def[symmetric]) |
|
2611 apply(simp add: fresh_perm) |
|
2612 apply(clarify) |
|
2613 apply(rotate_tac 2) |
|
2614 apply(drule subsetD[OF **]) |
|
2615 apply(simp add: inter_eqvt supp_eqvt permute_self) |
|
2616 done |
|
2617 ultimately have "(\<forall>b \<in> bs. q \<bullet> b = p \<bullet> b) \<and> supp q \<subseteq> bs \<union> (p \<bullet> bs)" using * by auto |
|
2618 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 |
|
2619 qed |
|
2620 |
|
2621 lemma list_renaming_perm: |
|
2622 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)" |
|
2623 proof (induct bs) |
|
2624 case (Cons a bs) |
|
2625 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 |
|
2626 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)" |
|
2627 by (blast) |
|
2628 { assume 1: "a \<in> set bs" |
|
2629 have "q \<bullet> a = p \<bullet> a" using * 1 by (induct bs) (auto) |
|
2630 then have "\<forall>b \<in> set (a # bs). q \<bullet> b = p \<bullet> b" using * by simp |
|
2631 moreover |
|
2632 have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" using ** by (auto simp add: insert_eqvt) |
|
2633 ultimately |
|
2634 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 |
|
2635 } |
|
2636 moreover |
|
2637 { assume 2: "a \<notin> set bs" |
|
2638 def q' \<equiv> "((q \<bullet> a) \<rightleftharpoons> (p \<bullet> a)) + q" |
|
2639 have "\<forall>b \<in> set (a # bs). q' \<bullet> b = p \<bullet> b" |
|
2640 unfolding q'_def using 2 * `a \<notin> set bs` by (auto simp add: swap_atom) |
|
2641 moreover |
|
2642 { have "{q \<bullet> a, p \<bullet> a} \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" |
|
2643 using ** |
|
2644 apply (auto simp add: supp_perm insert_eqvt) |
|
2645 apply (subgoal_tac "q \<bullet> a \<in> set bs \<union> p \<bullet> set bs") |
|
2646 apply(auto)[1] |
|
2647 apply(subgoal_tac "q \<bullet> a \<in> {a. q \<bullet> a \<noteq> a}") |
|
2648 apply(blast) |
|
2649 apply(simp) |
|
2650 done |
|
2651 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) |
|
2652 moreover |
|
2653 have "supp q \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" |
|
2654 using ** by (auto simp add: insert_eqvt) |
|
2655 ultimately |
|
2656 have "supp q' \<subseteq> set (a # bs) \<union> p \<bullet> (set (a # bs))" |
|
2657 unfolding q'_def using supp_plus_perm by blast |
|
2658 } |
|
2659 ultimately |
|
2660 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 |
|
2661 } |
|
2662 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))" |
|
2663 by blast |
|
2664 next |
|
2665 case Nil |
|
2666 have "(\<forall>b \<in> set []. 0 \<bullet> b = p \<bullet> b) \<and> supp (0::perm) \<subseteq> set [] \<union> p \<bullet> set []" |
|
2667 by (simp add: supp_zero_perm) |
|
2668 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 |
|
2669 qed |
|
2670 |
|
2671 |
|
2672 section {* Concrete Atoms Types *} |
|
2673 |
|
2674 text {* |
|
2675 Class @{text at_base} allows types containing multiple sorts of atoms. |
|
2676 Class @{text at} only allows types with a single sort. |
|
2677 *} |
|
2678 |
|
2679 class at_base = pt + |
|
2680 fixes atom :: "'a \<Rightarrow> atom" |
|
2681 assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b" |
|
2682 assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)" |
|
2683 |
|
2684 declare atom_eqvt[eqvt] |
|
2685 |
|
2686 class at = at_base + |
|
2687 assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)" |
|
2688 |
|
2689 lemma sort_ineq [simp]: |
|
2690 assumes "sort_of (atom a) \<noteq> sort_of (atom b)" |
|
2691 shows "atom a \<noteq> atom b" |
|
2692 using assms by metis |
|
2693 |
|
2694 lemma supp_at_base: |
|
2695 fixes a::"'a::at_base" |
|
2696 shows "supp a = {atom a}" |
|
2697 by (simp add: supp_atom [symmetric] supp_def atom_eqvt) |
|
2698 |
|
2699 lemma fresh_at_base: |
|
2700 shows "sort_of a \<noteq> sort_of (atom b) \<Longrightarrow> a \<sharp> b" |
|
2701 and "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b" |
|
2702 unfolding fresh_def |
|
2703 apply(simp_all add: supp_at_base) |
|
2704 apply(metis) |
|
2705 done |
|
2706 |
|
2707 lemma fresh_atom_at_base: |
|
2708 fixes b::"'a::at_base" |
|
2709 shows "a \<sharp> atom b \<longleftrightarrow> a \<sharp> b" |
|
2710 by (simp add: fresh_def supp_at_base supp_atom) |
|
2711 |
|
2712 lemma fresh_star_atom_at_base: |
|
2713 fixes b::"'a::at_base" |
|
2714 shows "as \<sharp>* atom b \<longleftrightarrow> as \<sharp>* b" |
|
2715 by (simp add: fresh_star_def fresh_atom_at_base) |
|
2716 |
|
2717 instance at_base < fs |
|
2718 proof qed (simp add: supp_at_base) |
|
2719 |
|
2720 lemma at_base_infinite [simp]: |
|
2721 shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U") |
|
2722 proof |
|
2723 obtain a :: 'a where "True" by auto |
|
2724 assume "finite ?U" |
|
2725 hence "finite (atom ` ?U)" |
|
2726 by (rule finite_imageI) |
|
2727 then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)" |
|
2728 by (rule obtain_atom) |
|
2729 from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)" |
|
2730 unfolding atom_eqvt [symmetric] |
|
2731 by (simp add: swap_atom) |
|
2732 hence "b \<in> atom ` ?U" by simp |
|
2733 with b(1) show "False" by simp |
|
2734 qed |
|
2735 |
|
2736 lemma swap_at_base_simps [simp]: |
|
2737 fixes x y::"'a::at_base" |
|
2738 shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y" |
|
2739 and "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x" |
|
2740 and "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" |
|
2741 unfolding atom_eq_iff [symmetric] |
|
2742 unfolding atom_eqvt [symmetric] |
|
2743 by simp_all |
|
2744 |
|
2745 lemma obtain_at_base: |
|
2746 assumes X: "finite X" |
|
2747 obtains a::"'a::at_base" where "atom a \<notin> X" |
|
2748 proof - |
|
2749 have "inj (atom :: 'a \<Rightarrow> atom)" |
|
2750 by (simp add: inj_on_def) |
|
2751 with X have "finite (atom -` X :: 'a set)" |
|
2752 by (rule finite_vimageI) |
|
2753 with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)" |
|
2754 by auto |
|
2755 then obtain a :: 'a where "atom a \<notin> X" |
|
2756 by auto |
|
2757 thus ?thesis .. |
|
2758 qed |
|
2759 |
|
2760 lemma obtain_fresh': |
|
2761 assumes fin: "finite (supp x)" |
|
2762 obtains a::"'a::at_base" where "atom a \<sharp> x" |
|
2763 using obtain_at_base[where X="supp x"] |
|
2764 by (auto simp add: fresh_def fin) |
|
2765 |
|
2766 lemma obtain_fresh: |
|
2767 fixes x::"'b::fs" |
|
2768 obtains a::"'a::at_base" where "atom a \<sharp> x" |
|
2769 by (rule obtain_fresh') (auto simp add: finite_supp) |
|
2770 |
|
2771 lemma supp_finite_set_at_base: |
|
2772 assumes a: "finite S" |
|
2773 shows "supp S = atom ` S" |
|
2774 apply(simp add: supp_of_finite_sets[OF a]) |
|
2775 apply(simp add: supp_at_base) |
|
2776 apply(auto) |
|
2777 done |
|
2778 |
|
2779 (* FIXME |
|
2780 lemma supp_cofinite_set_at_base: |
|
2781 assumes a: "finite (UNIV - S)" |
|
2782 shows "supp S = atom ` (UNIV - S)" |
|
2783 apply(rule finite_supp_unique) |
|
2784 *) |
|
2785 |
|
2786 lemma fresh_finite_set_at_base: |
|
2787 fixes a::"'a::at_base" |
|
2788 assumes a: "finite S" |
|
2789 shows "atom a \<sharp> S \<longleftrightarrow> a \<notin> S" |
|
2790 unfolding fresh_def |
|
2791 apply(simp add: supp_finite_set_at_base[OF a]) |
|
2792 apply(subst inj_image_mem_iff) |
|
2793 apply(simp add: inj_on_def) |
|
2794 apply(simp) |
|
2795 done |
|
2796 |
|
2797 lemma fresh_at_base_permute_iff [simp]: |
|
2798 fixes a::"'a::at_base" |
|
2799 shows "atom (p \<bullet> a) \<sharp> p \<bullet> x \<longleftrightarrow> atom a \<sharp> x" |
|
2800 unfolding atom_eqvt[symmetric] |
|
2801 by (simp add: fresh_permute_iff) |
|
2802 |
|
2803 |
|
2804 section {* Infrastructure for concrete atom types *} |
|
2805 |
|
2806 definition |
|
2807 flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')") |
|
2808 where |
|
2809 "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)" |
|
2810 |
|
2811 lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0" |
|
2812 unfolding flip_def by (rule swap_self) |
|
2813 |
|
2814 lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)" |
|
2815 unfolding flip_def by (rule swap_commute) |
|
2816 |
|
2817 lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)" |
|
2818 unfolding flip_def by (rule minus_swap) |
|
2819 |
|
2820 lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0" |
|
2821 unfolding flip_def by (rule swap_cancel) |
|
2822 |
|
2823 lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x" |
|
2824 unfolding permute_plus [symmetric] add_flip_cancel by simp |
|
2825 |
|
2826 lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x" |
|
2827 by (simp add: flip_commute) |
|
2828 |
|
2829 lemma flip_eqvt [eqvt]: |
|
2830 fixes a b c::"'a::at_base" |
|
2831 shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)" |
|
2832 unfolding flip_def |
|
2833 by (simp add: swap_eqvt atom_eqvt) |
|
2834 |
|
2835 lemma flip_at_base_simps [simp]: |
|
2836 shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b" |
|
2837 and "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a" |
|
2838 and "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c" |
|
2839 and "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x" |
|
2840 unfolding flip_def |
|
2841 unfolding atom_eq_iff [symmetric] |
|
2842 unfolding atom_eqvt [symmetric] |
|
2843 by simp_all |
|
2844 |
|
2845 text {* the following two lemmas do not hold for at_base, |
|
2846 only for single sort atoms from at *} |
|
2847 |
|
2848 lemma permute_flip_at: |
|
2849 fixes a b c::"'a::at" |
|
2850 shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)" |
|
2851 unfolding flip_def |
|
2852 apply (rule atom_eq_iff [THEN iffD1]) |
|
2853 apply (subst atom_eqvt [symmetric]) |
|
2854 apply (simp add: swap_atom) |
|
2855 done |
|
2856 |
|
2857 lemma flip_at_simps [simp]: |
|
2858 fixes a b::"'a::at" |
|
2859 shows "(a \<leftrightarrow> b) \<bullet> a = b" |
|
2860 and "(a \<leftrightarrow> b) \<bullet> b = a" |
|
2861 unfolding permute_flip_at by simp_all |
|
2862 |
|
2863 lemma flip_fresh_fresh: |
|
2864 fixes a b::"'a::at_base" |
|
2865 assumes "atom a \<sharp> x" "atom b \<sharp> x" |
|
2866 shows "(a \<leftrightarrow> b) \<bullet> x = x" |
|
2867 using assms |
|
2868 by (simp add: flip_def swap_fresh_fresh) |
|
2869 |
|
2870 |
|
2871 |
|
2872 subsection {* Syntax for coercing at-elements to the atom-type *} |
|
2873 |
|
2874 syntax |
|
2875 "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3) |
|
2876 |
|
2877 translations |
|
2878 "_atom_constrain a t" => "CONST atom (_constrain a t)" |
|
2879 |
|
2880 |
|
2881 subsection {* A lemma for proving instances of class @{text at}. *} |
|
2882 |
|
2883 setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *} |
|
2884 setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *} |
|
2885 |
|
2886 text {* |
|
2887 New atom types are defined as subtypes of @{typ atom}. |
|
2888 *} |
|
2889 |
|
2890 lemma exists_eq_simple_sort: |
|
2891 shows "\<exists>a. a \<in> {a. sort_of a = s}" |
|
2892 by (rule_tac x="Atom s 0" in exI, simp) |
|
2893 |
|
2894 lemma exists_eq_sort: |
|
2895 shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}" |
|
2896 by (rule_tac x="Atom (sort_fun x) y" in exI, simp) |
|
2897 |
|
2898 lemma at_base_class: |
|
2899 fixes sort_fun :: "'b \<Rightarrow> atom_sort" |
|
2900 fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" |
|
2901 assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}" |
|
2902 assumes atom_def: "\<And>a. atom a = Rep a" |
|
2903 assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" |
|
2904 shows "OFCLASS('a, at_base_class)" |
|
2905 proof |
|
2906 interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type) |
|
2907 have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp |
|
2908 fix a b :: 'a and p p1 p2 :: perm |
|
2909 show "0 \<bullet> a = a" |
|
2910 unfolding permute_def by (simp add: Rep_inverse) |
|
2911 show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" |
|
2912 unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) |
|
2913 show "atom a = atom b \<longleftrightarrow> a = b" |
|
2914 unfolding atom_def by (simp add: Rep_inject) |
|
2915 show "p \<bullet> atom a = atom (p \<bullet> a)" |
|
2916 unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) |
|
2917 qed |
|
2918 |
|
2919 (* |
|
2920 lemma at_class: |
|
2921 fixes s :: atom_sort |
|
2922 fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" |
|
2923 assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}" |
|
2924 assumes atom_def: "\<And>a. atom a = Rep a" |
|
2925 assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" |
|
2926 shows "OFCLASS('a, at_class)" |
|
2927 proof |
|
2928 interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type) |
|
2929 have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def) |
|
2930 fix a b :: 'a and p p1 p2 :: perm |
|
2931 show "0 \<bullet> a = a" |
|
2932 unfolding permute_def by (simp add: Rep_inverse) |
|
2933 show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" |
|
2934 unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) |
|
2935 show "sort_of (atom a) = sort_of (atom b)" |
|
2936 unfolding atom_def by (simp add: sort_of_Rep) |
|
2937 show "atom a = atom b \<longleftrightarrow> a = b" |
|
2938 unfolding atom_def by (simp add: Rep_inject) |
|
2939 show "p \<bullet> atom a = atom (p \<bullet> a)" |
|
2940 unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) |
|
2941 qed |
|
2942 *) |
|
2943 |
|
2944 lemma at_class: |
|
2945 fixes s :: atom_sort |
|
2946 fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" |
|
2947 assumes type: "type_definition Rep Abs {a. sort_of a = s}" |
|
2948 assumes atom_def: "\<And>a. atom a = Rep a" |
|
2949 assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)" |
|
2950 shows "OFCLASS('a, at_class)" |
|
2951 proof |
|
2952 interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type) |
|
2953 have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def) |
|
2954 fix a b :: 'a and p p1 p2 :: perm |
|
2955 show "0 \<bullet> a = a" |
|
2956 unfolding permute_def by (simp add: Rep_inverse) |
|
2957 show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a" |
|
2958 unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) |
|
2959 show "sort_of (atom a) = sort_of (atom b)" |
|
2960 unfolding atom_def by (simp add: sort_of_Rep) |
|
2961 show "atom a = atom b \<longleftrightarrow> a = b" |
|
2962 unfolding atom_def by (simp add: Rep_inject) |
|
2963 show "p \<bullet> atom a = atom (p \<bullet> a)" |
|
2964 unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) |
|
2965 qed |
|
2966 |
|
2967 lemma at_class_sort: |
|
2968 fixes s :: atom_sort |
|
2969 fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a" |
|
2970 fixes a::"'a" |
|
2971 assumes type: "type_definition Rep Abs {a. sort_of a = s}" |
|
2972 assumes atom_def: "\<And>a. atom a = Rep a" |
|
2973 shows "sort_of (atom a) = s" |
|
2974 using atom_def type |
|
2975 unfolding type_definition_def by simp |
|
2976 |
|
2977 |
|
2978 setup {* Sign.add_const_constraint |
|
2979 (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *} |
|
2980 setup {* Sign.add_const_constraint |
|
2981 (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *} |
|
2982 |
|
2983 section {* The freshness lemma according to Andy Pitts *} |
|
2984 |
|
2985 lemma freshness_lemma: |
|
2986 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
2987 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
|
2988 shows "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" |
|
2989 proof - |
|
2990 from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b" |
|
2991 by (auto simp add: fresh_Pair) |
|
2992 show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" |
|
2993 proof (intro exI allI impI) |
|
2994 fix a :: 'a |
|
2995 assume a3: "atom a \<sharp> h" |
|
2996 show "h a = h b" |
|
2997 proof (cases "a = b") |
|
2998 assume "a = b" |
|
2999 thus "h a = h b" by simp |
|
3000 next |
|
3001 assume "a \<noteq> b" |
|
3002 hence "atom a \<sharp> b" by (simp add: fresh_at_base) |
|
3003 with a3 have "atom a \<sharp> h b" |
|
3004 by (rule fresh_fun_app) |
|
3005 with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)" |
|
3006 by (rule swap_fresh_fresh) |
|
3007 from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h" |
|
3008 by (rule swap_fresh_fresh) |
|
3009 from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp |
|
3010 also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)" |
|
3011 by (rule permute_fun_app_eq) |
|
3012 also have "\<dots> = h a" |
|
3013 using d2 by simp |
|
3014 finally show "h a = h b" by simp |
|
3015 qed |
|
3016 qed |
|
3017 qed |
|
3018 |
|
3019 lemma freshness_lemma_unique: |
|
3020 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
3021 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
|
3022 shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" |
|
3023 proof (rule ex_ex1I) |
|
3024 from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" |
|
3025 by (rule freshness_lemma) |
|
3026 next |
|
3027 fix x y |
|
3028 assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" |
|
3029 assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y" |
|
3030 from a x y show "x = y" |
|
3031 by (auto simp add: fresh_Pair) |
|
3032 qed |
|
3033 |
|
3034 text {* packaging the freshness lemma into a function *} |
|
3035 |
|
3036 definition |
|
3037 fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b" |
|
3038 where |
|
3039 "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)" |
|
3040 |
|
3041 lemma fresh_fun_apply: |
|
3042 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
3043 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
|
3044 assumes b: "atom a \<sharp> h" |
|
3045 shows "fresh_fun h = h a" |
|
3046 unfolding fresh_fun_def |
|
3047 proof (rule the_equality) |
|
3048 show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a" |
|
3049 proof (intro strip) |
|
3050 fix a':: 'a |
|
3051 assume c: "atom a' \<sharp> h" |
|
3052 from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma) |
|
3053 with b c show "h a' = h a" by auto |
|
3054 qed |
|
3055 next |
|
3056 fix fr :: 'b |
|
3057 assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr" |
|
3058 with b show "fr = h a" by auto |
|
3059 qed |
|
3060 |
|
3061 lemma fresh_fun_apply': |
|
3062 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
3063 assumes a: "atom a \<sharp> h" "atom a \<sharp> h a" |
|
3064 shows "fresh_fun h = h a" |
|
3065 apply (rule fresh_fun_apply) |
|
3066 apply (auto simp add: fresh_Pair intro: a) |
|
3067 done |
|
3068 |
|
3069 lemma fresh_fun_eqvt: |
|
3070 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
3071 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
|
3072 shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)" |
|
3073 using a |
|
3074 apply (clarsimp simp add: fresh_Pair) |
|
3075 apply (subst fresh_fun_apply', assumption+) |
|
3076 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
|
3077 apply (drule fresh_permute_iff [where p=p, THEN iffD2]) |
|
3078 apply (simp only: atom_eqvt permute_fun_app_eq [where f=h]) |
|
3079 apply (erule (1) fresh_fun_apply' [symmetric]) |
|
3080 done |
|
3081 |
|
3082 lemma fresh_fun_supports: |
|
3083 fixes h :: "'a::at \<Rightarrow> 'b::pt" |
|
3084 assumes a: "\<exists>a. atom a \<sharp> (h, h a)" |
|
3085 shows "(supp h) supports (fresh_fun h)" |
|
3086 apply (simp add: supports_def fresh_def [symmetric]) |
|
3087 apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh) |
|
3088 done |
|
3089 |
|
3090 notation fresh_fun (binder "FRESH " 10) |
|
3091 |
|
3092 lemma FRESH_f_iff: |
|
3093 fixes P :: "'a::at \<Rightarrow> 'b::pure" |
|
3094 fixes f :: "'b \<Rightarrow> 'c::pure" |
|
3095 assumes P: "finite (supp P)" |
|
3096 shows "(FRESH x. f (P x)) = f (FRESH x. P x)" |
|
3097 proof - |
|
3098 obtain a::'a where "atom a \<sharp> P" using P by (rule obtain_fresh') |
|
3099 show "(FRESH x. f (P x)) = f (FRESH x. P x)" |
|
3100 apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) |
|
3101 apply (cut_tac `atom a \<sharp> P`) |
|
3102 apply (simp add: fresh_conv_MOST) |
|
3103 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
|
3104 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
|
3105 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
|
3106 apply (rule refl) |
|
3107 done |
|
3108 qed |
|
3109 |
|
3110 lemma FRESH_binop_iff: |
|
3111 fixes P :: "'a::at \<Rightarrow> 'b::pure" |
|
3112 fixes Q :: "'a::at \<Rightarrow> 'c::pure" |
|
3113 fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure" |
|
3114 assumes P: "finite (supp P)" |
|
3115 and Q: "finite (supp Q)" |
|
3116 shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)" |
|
3117 proof - |
|
3118 from assms have "finite (supp (P, Q))" by (simp add: supp_Pair) |
|
3119 then obtain a::'a where "atom a \<sharp> (P, Q)" by (rule obtain_fresh') |
|
3120 then have "atom a \<sharp> P" and "atom a \<sharp> Q" by (simp_all add: fresh_Pair) |
|
3121 show ?thesis |
|
3122 apply (subst fresh_fun_apply' [where a=a, OF _ pure_fresh]) |
|
3123 apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`) |
|
3124 apply (simp add: fresh_conv_MOST) |
|
3125 apply (elim MOST_rev_mp, rule MOST_I, clarify) |
|
3126 apply (simp add: permute_fun_def permute_pure fun_eq_iff) |
|
3127 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> P` pure_fresh]) |
|
3128 apply (subst fresh_fun_apply' [where a=a, OF `atom a \<sharp> Q` pure_fresh]) |
|
3129 apply (rule refl) |
|
3130 done |
|
3131 qed |
|
3132 |
|
3133 lemma FRESH_conj_iff: |
|
3134 fixes P Q :: "'a::at \<Rightarrow> bool" |
|
3135 assumes P: "finite (supp P)" and Q: "finite (supp Q)" |
|
3136 shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)" |
|
3137 using P Q by (rule FRESH_binop_iff) |
|
3138 |
|
3139 lemma FRESH_disj_iff: |
|
3140 fixes P Q :: "'a::at \<Rightarrow> bool" |
|
3141 assumes P: "finite (supp P)" and Q: "finite (supp Q)" |
|
3142 shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)" |
|
3143 using P Q by (rule FRESH_binop_iff) |
|
3144 |
|
3145 |
|
3146 section {* Library functions for the nominal infrastructure *} |
|
3147 |
|
3148 use "nominal_library.ML" |
|
3149 |
|
3150 |
|
3151 section {* Automation for creating concrete atom types *} |
|
3152 |
|
3153 text {* at the moment only single-sort concrete atoms are supported *} |
|
3154 |
|
3155 use "nominal_atoms.ML" |
|
3156 |
|
3157 |
|
3158 section {* automatic equivariance procedure for inductive definitions *} |
|
3159 |
|
3160 use "nominal_eqvt.ML" |
|
3161 |
|
3162 |
|
3163 |
|
3164 |
|
3165 end |