| author | wenzelm | 
| Sat, 07 Jan 2006 12:26:33 +0100 | |
| changeset 18608 | 9cdcc2a5c8b3 | 
| parent 18600 | 20ad06db427b | 
| child 18627 | f0acb66858b4 | 
| permissions | -rw-r--r-- | 
| 17870 | 1  | 
(* $Id$ *)  | 
2  | 
||
3  | 
theory nominal  | 
|
4  | 
imports Main  | 
|
| 18068 | 5  | 
uses  | 
6  | 
  ("nominal_atoms.ML")
 | 
|
7  | 
  ("nominal_package.ML")
 | 
|
| 18264 | 8  | 
  ("nominal_induct.ML") 
 | 
| 18068 | 9  | 
  ("nominal_permeq.ML")
 | 
| 17870 | 10  | 
begin  | 
11  | 
||
12  | 
ML {* reset NameSpace.unique_names; *}
 | 
|
13  | 
||
14  | 
section {* Permutations *}
 | 
|
15  | 
(*======================*)  | 
|
16  | 
||
17  | 
types  | 
|
18  | 
  'x prm = "('x \<times> 'x) list"
 | 
|
19  | 
||
20  | 
(* polymorphic operations for permutation and swapping*)  | 
|
21  | 
consts  | 
|
| 18491 | 22  | 
perm :: "'x prm \<Rightarrow> 'a \<Rightarrow> 'a" (infixr "\<bullet>" 80)  | 
| 17870 | 23  | 
  swap :: "('x \<times> 'x) \<Rightarrow> 'x \<Rightarrow> 'x"
 | 
24  | 
||
25  | 
(* permutation on sets *)  | 
|
26  | 
defs (overloaded)  | 
|
27  | 
  perm_set_def:  "pi\<bullet>(X::'a set) \<equiv> {pi\<bullet>a | a. a\<in>X}"
 | 
|
28  | 
||
| 18264 | 29  | 
lemma perm_union:  | 
30  | 
shows "pi \<bullet> (X \<union> Y) = (pi \<bullet> X) \<union> (pi \<bullet> Y)"  | 
|
31  | 
by (auto simp add: perm_set_def)  | 
|
32  | 
||
| 17870 | 33  | 
(* permutation on units and products *)  | 
34  | 
primrec (perm_unit)  | 
|
35  | 
"pi\<bullet>() = ()"  | 
|
36  | 
||
37  | 
primrec (perm_prod)  | 
|
38  | 
"pi\<bullet>(a,b) = (pi\<bullet>a,pi\<bullet>b)"  | 
|
39  | 
||
40  | 
lemma perm_fst:  | 
|
41  | 
"pi\<bullet>(fst x) = fst (pi\<bullet>x)"  | 
|
42  | 
by (cases x, simp)  | 
|
43  | 
||
44  | 
lemma perm_snd:  | 
|
45  | 
"pi\<bullet>(snd x) = snd (pi\<bullet>x)"  | 
|
46  | 
by (cases x, simp)  | 
|
47  | 
||
48  | 
(* permutation on lists *)  | 
|
49  | 
primrec (perm_list)  | 
|
50  | 
perm_nil_def: "pi\<bullet>[] = []"  | 
|
51  | 
perm_cons_def: "pi\<bullet>(x#xs) = (pi\<bullet>x)#(pi\<bullet>xs)"  | 
|
52  | 
||
53  | 
lemma perm_append:  | 
|
54  | 
fixes pi :: "'x prm"  | 
|
55  | 
and l1 :: "'a list"  | 
|
56  | 
and l2 :: "'a list"  | 
|
57  | 
shows "pi\<bullet>(l1@l2) = (pi\<bullet>l1)@(pi\<bullet>l2)"  | 
|
58  | 
by (induct l1, auto)  | 
|
59  | 
||
60  | 
lemma perm_rev:  | 
|
61  | 
fixes pi :: "'x prm"  | 
|
62  | 
and l :: "'a list"  | 
|
63  | 
shows "pi\<bullet>(rev l) = rev (pi\<bullet>l)"  | 
|
64  | 
by (induct l, simp_all add: perm_append)  | 
|
65  | 
||
66  | 
(* permutation on functions *)  | 
|
67  | 
defs (overloaded)  | 
|
68  | 
perm_fun_def: "pi\<bullet>(f::'a\<Rightarrow>'b) \<equiv> (\<lambda>x. pi\<bullet>f((rev pi)\<bullet>x))"  | 
|
69  | 
||
70  | 
(* permutation on bools *)  | 
|
71  | 
primrec (perm_bool)  | 
|
72  | 
perm_true_def: "pi\<bullet>True = True"  | 
|
73  | 
perm_false_def: "pi\<bullet>False = False"  | 
|
74  | 
||
| 18264 | 75  | 
lemma perm_bool:  | 
76  | 
shows "pi\<bullet>(b::bool) = b"  | 
|
77  | 
by (cases "b", auto)  | 
|
78  | 
||
| 17870 | 79  | 
(* permutation on options *)  | 
80  | 
primrec (perm_option)  | 
|
81  | 
perm_some_def: "pi\<bullet>Some(x) = Some(pi\<bullet>x)"  | 
|
82  | 
perm_none_def: "pi\<bullet>None = None"  | 
|
83  | 
||
84  | 
(* a "private" copy of the option type used in the abstraction function *)  | 
|
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
85  | 
datatype 'a noption = nSome 'a | nNone  | 
| 17870 | 86  | 
|
87  | 
primrec (perm_noption)  | 
|
| 18600 | 88  | 
perm_nSome_def: "pi\<bullet>nSome(x) = nSome(pi\<bullet>x)"  | 
89  | 
perm_nNone_def: "pi\<bullet>nNone = nNone"  | 
|
90  | 
||
91  | 
(* a "private" copy of the product type used in the nominal induct method *)  | 
|
92  | 
datatype ('a,'b) nprod = nPair 'a 'b
 | 
|
93  | 
||
94  | 
primrec (perm_nprod)  | 
|
95  | 
perm_nProd_def: "pi\<bullet>(nPair x1 x2) = nPair (pi\<bullet>x1) (pi\<bullet>x2)"  | 
|
| 17870 | 96  | 
|
97  | 
(* permutation on characters (used in strings) *)  | 
|
98  | 
defs (overloaded)  | 
|
99  | 
perm_char_def: "pi\<bullet>(s::char) \<equiv> s"  | 
|
100  | 
||
101  | 
(* permutation on ints *)  | 
|
102  | 
defs (overloaded)  | 
|
103  | 
perm_int_def: "pi\<bullet>(i::int) \<equiv> i"  | 
|
104  | 
||
105  | 
(* permutation on nats *)  | 
|
106  | 
defs (overloaded)  | 
|
107  | 
perm_nat_def: "pi\<bullet>(i::nat) \<equiv> i"  | 
|
108  | 
||
109  | 
section {* permutation equality *}
 | 
|
110  | 
(*==============================*)  | 
|
111  | 
||
112  | 
constdefs  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
113  | 
  prm_eq :: "'x prm \<Rightarrow> 'x prm \<Rightarrow> bool"  (" _ \<triangleq> _ " [80,80] 80)
 | 
| 
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
114  | 
"pi1 \<triangleq> pi2 \<equiv> \<forall>a::'x. pi1\<bullet>a = pi2\<bullet>a"  | 
| 17870 | 115  | 
|
116  | 
section {* Support, Freshness and Supports*}
 | 
|
117  | 
(*========================================*)  | 
|
118  | 
constdefs  | 
|
119  | 
   supp :: "'a \<Rightarrow> ('x set)"  
 | 
|
120  | 
   "supp x \<equiv> {a . (infinite {b . [(a,b)]\<bullet>x \<noteq> x})}"
 | 
|
121  | 
||
| 17871 | 122  | 
   fresh :: "'x \<Rightarrow> 'a \<Rightarrow> bool" ("_ \<sharp> _" [80,80] 80)
 | 
| 17870 | 123  | 
"a \<sharp> x \<equiv> a \<notin> supp x"  | 
124  | 
||
125  | 
supports :: "'x set \<Rightarrow> 'a \<Rightarrow> bool" (infixl 80)  | 
|
126  | 
"S supports x \<equiv> \<forall>a b. (a\<notin>S \<and> b\<notin>S \<longrightarrow> [(a,b)]\<bullet>x=x)"  | 
|
127  | 
||
128  | 
lemma supp_fresh_iff:  | 
|
129  | 
fixes x :: "'a"  | 
|
130  | 
  shows "(supp x) = {a::'x. \<not>a\<sharp>x}"
 | 
|
131  | 
apply(simp add: fresh_def)  | 
|
132  | 
done  | 
|
133  | 
||
134  | 
lemma supp_unit:  | 
|
135  | 
  shows "supp () = {}"
 | 
|
136  | 
by (simp add: supp_def)  | 
|
137  | 
||
| 18264 | 138  | 
lemma supp_set_empty:  | 
139  | 
  shows "supp {} = {}"
 | 
|
140  | 
by (force simp add: supp_def perm_set_def)  | 
|
141  | 
||
142  | 
lemma supp_singleton:  | 
|
143  | 
  shows "supp {x} = supp x"
 | 
|
144  | 
by (force simp add: supp_def perm_set_def)  | 
|
145  | 
||
| 17870 | 146  | 
lemma supp_prod:  | 
147  | 
fixes x :: "'a"  | 
|
148  | 
and y :: "'b"  | 
|
149  | 
shows "(supp (x,y)) = (supp x)\<union>(supp y)"  | 
|
150  | 
by (force simp add: supp_def Collect_imp_eq Collect_neg_eq)  | 
|
151  | 
||
| 18600 | 152  | 
lemma supp_nprod:  | 
153  | 
fixes x :: "'a"  | 
|
154  | 
and y :: "'b"  | 
|
155  | 
shows "(supp (nPair x y)) = (supp x)\<union>(supp y)"  | 
|
156  | 
by (force simp add: supp_def Collect_imp_eq Collect_neg_eq)  | 
|
157  | 
||
| 17870 | 158  | 
lemma supp_list_nil:  | 
159  | 
  shows "supp [] = {}"
 | 
|
160  | 
apply(simp add: supp_def)  | 
|
161  | 
done  | 
|
162  | 
||
163  | 
lemma supp_list_cons:  | 
|
164  | 
fixes x :: "'a"  | 
|
165  | 
and xs :: "'a list"  | 
|
166  | 
shows "supp (x#xs) = (supp x)\<union>(supp xs)"  | 
|
167  | 
apply(auto simp add: supp_def Collect_imp_eq Collect_neg_eq)  | 
|
168  | 
done  | 
|
169  | 
||
170  | 
lemma supp_list_append:  | 
|
171  | 
fixes xs :: "'a list"  | 
|
172  | 
and ys :: "'a list"  | 
|
173  | 
shows "supp (xs@ys) = (supp xs)\<union>(supp ys)"  | 
|
174  | 
by (induct xs, auto simp add: supp_list_nil supp_list_cons)  | 
|
175  | 
||
176  | 
lemma supp_list_rev:  | 
|
177  | 
fixes xs :: "'a list"  | 
|
178  | 
shows "supp (rev xs) = (supp xs)"  | 
|
179  | 
by (induct xs, auto simp add: supp_list_append supp_list_cons supp_list_nil)  | 
|
180  | 
||
181  | 
lemma supp_bool:  | 
|
182  | 
fixes x :: "bool"  | 
|
183  | 
  shows "supp (x) = {}"
 | 
|
184  | 
apply(case_tac "x")  | 
|
185  | 
apply(simp_all add: supp_def)  | 
|
186  | 
done  | 
|
187  | 
||
188  | 
lemma supp_some:  | 
|
189  | 
fixes x :: "'a"  | 
|
190  | 
shows "supp (Some x) = (supp x)"  | 
|
191  | 
apply(simp add: supp_def)  | 
|
192  | 
done  | 
|
193  | 
||
194  | 
lemma supp_none:  | 
|
195  | 
fixes x :: "'a"  | 
|
196  | 
  shows "supp (None) = {}"
 | 
|
197  | 
apply(simp add: supp_def)  | 
|
198  | 
done  | 
|
199  | 
||
200  | 
lemma supp_int:  | 
|
201  | 
fixes i::"int"  | 
|
202  | 
  shows "supp (i) = {}"
 | 
|
203  | 
apply(simp add: supp_def perm_int_def)  | 
|
204  | 
done  | 
|
205  | 
||
| 18264 | 206  | 
lemma fresh_set_empty:  | 
207  | 
  shows "a\<sharp>{}"
 | 
|
208  | 
by (simp add: fresh_def supp_set_empty)  | 
|
209  | 
||
| 18578 | 210  | 
lemma fresh_singleton:  | 
211  | 
  shows "a\<sharp>{x} = a\<sharp>x"
 | 
|
212  | 
by (simp add: fresh_def supp_singleton)  | 
|
213  | 
||
| 17870 | 214  | 
lemma fresh_prod:  | 
215  | 
fixes a :: "'x"  | 
|
216  | 
and x :: "'a"  | 
|
217  | 
and y :: "'b"  | 
|
218  | 
shows "a\<sharp>(x,y) = (a\<sharp>x \<and> a\<sharp>y)"  | 
|
219  | 
by (simp add: fresh_def supp_prod)  | 
|
220  | 
||
221  | 
lemma fresh_list_nil:  | 
|
222  | 
fixes a :: "'x"  | 
|
| 18264 | 223  | 
shows "a\<sharp>[]"  | 
| 17870 | 224  | 
by (simp add: fresh_def supp_list_nil)  | 
225  | 
||
226  | 
lemma fresh_list_cons:  | 
|
227  | 
fixes a :: "'x"  | 
|
228  | 
and x :: "'a"  | 
|
229  | 
and xs :: "'a list"  | 
|
230  | 
shows "a\<sharp>(x#xs) = (a\<sharp>x \<and> a\<sharp>xs)"  | 
|
231  | 
by (simp add: fresh_def supp_list_cons)  | 
|
232  | 
||
233  | 
lemma fresh_list_append:  | 
|
234  | 
fixes a :: "'x"  | 
|
235  | 
and xs :: "'a list"  | 
|
236  | 
and ys :: "'a list"  | 
|
237  | 
shows "a\<sharp>(xs@ys) = (a\<sharp>xs \<and> a\<sharp>ys)"  | 
|
238  | 
by (simp add: fresh_def supp_list_append)  | 
|
239  | 
||
240  | 
lemma fresh_list_rev:  | 
|
241  | 
fixes a :: "'x"  | 
|
242  | 
and xs :: "'a list"  | 
|
243  | 
shows "a\<sharp>(rev xs) = a\<sharp>xs"  | 
|
244  | 
by (simp add: fresh_def supp_list_rev)  | 
|
245  | 
||
246  | 
lemma fresh_none:  | 
|
247  | 
fixes a :: "'x"  | 
|
248  | 
shows "a\<sharp>None"  | 
|
249  | 
apply(simp add: fresh_def supp_none)  | 
|
250  | 
done  | 
|
251  | 
||
252  | 
lemma fresh_some:  | 
|
253  | 
fixes a :: "'x"  | 
|
254  | 
and x :: "'a"  | 
|
255  | 
shows "a\<sharp>(Some x) = a\<sharp>x"  | 
|
256  | 
apply(simp add: fresh_def supp_some)  | 
|
257  | 
done  | 
|
258  | 
||
| 
18294
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
259  | 
text {* Normalization of freshness results; cf.\ @{text nominal_induct} *}
 | 
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
260  | 
|
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
261  | 
lemma fresh_unit_elim: "(a\<sharp>() \<Longrightarrow> PROP C) == PROP C"  | 
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
262  | 
by (simp add: fresh_def supp_unit)  | 
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
263  | 
|
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
264  | 
lemma fresh_prod_elim: "(a\<sharp>(x,y) \<Longrightarrow> PROP C) == (a\<sharp>x \<Longrightarrow> a\<sharp>y \<Longrightarrow> PROP C)"  | 
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
265  | 
by rule (simp_all add: fresh_prod)  | 
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
266  | 
|
| 
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
267  | 
|
| 17870 | 268  | 
section {* Abstract Properties for Permutations and  Atoms *}
 | 
269  | 
(*=========================================================*)  | 
|
270  | 
||
271  | 
(* properties for being a permutation type *)  | 
|
272  | 
constdefs  | 
|
273  | 
  "pt TYPE('a) TYPE('x) \<equiv> 
 | 
|
274  | 
(\<forall>(x::'a). ([]::'x prm)\<bullet>x = x) \<and>  | 
|
275  | 
(\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). (pi1@pi2)\<bullet>x = pi1\<bullet>(pi2\<bullet>x)) \<and>  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
276  | 
(\<forall>(pi1::'x prm) (pi2::'x prm) (x::'a). pi1 \<triangleq> pi2 \<longrightarrow> pi1\<bullet>x = pi2\<bullet>x)"  | 
| 17870 | 277  | 
|
278  | 
(* properties for being an atom type *)  | 
|
279  | 
constdefs  | 
|
280  | 
  "at TYPE('x) \<equiv> 
 | 
|
281  | 
(\<forall>(x::'x). ([]::'x prm)\<bullet>x = x) \<and>  | 
|
282  | 
(\<forall>(a::'x) (b::'x) (pi::'x prm) (x::'x). ((a,b)#(pi::'x prm))\<bullet>x = swap (a,b) (pi\<bullet>x)) \<and>  | 
|
283  | 
(\<forall>(a::'x) (b::'x) (c::'x). swap (a,b) c = (if a=c then b else (if b=c then a else c))) \<and>  | 
|
284  | 
(infinite (UNIV::'x set))"  | 
|
285  | 
||
286  | 
(* property of two atom-types being disjoint *)  | 
|
287  | 
constdefs  | 
|
288  | 
  "disjoint TYPE('x) TYPE('y) \<equiv> 
 | 
|
289  | 
(\<forall>(pi::'x prm)(x::'y). pi\<bullet>x = x) \<and>  | 
|
290  | 
(\<forall>(pi::'y prm)(x::'x). pi\<bullet>x = x)"  | 
|
291  | 
||
292  | 
(* composition property of two permutation on a type 'a *)  | 
|
293  | 
constdefs  | 
|
294  | 
  "cp TYPE ('a) TYPE('x) TYPE('y) \<equiv> 
 | 
|
295  | 
(\<forall>(pi2::'y prm) (pi1::'x prm) (x::'a) . pi1\<bullet>(pi2\<bullet>x) = (pi1\<bullet>pi2)\<bullet>(pi1\<bullet>x))"  | 
|
296  | 
||
297  | 
(* property of having finite support *)  | 
|
298  | 
constdefs  | 
|
299  | 
  "fs TYPE('a) TYPE('x) \<equiv> \<forall>(x::'a). finite ((supp x)::'x set)"
 | 
|
300  | 
||
301  | 
section {* Lemmas about the atom-type properties*}
 | 
|
302  | 
(*==============================================*)  | 
|
303  | 
||
304  | 
lemma at1:  | 
|
305  | 
fixes x::"'x"  | 
|
306  | 
  assumes a: "at TYPE('x)"
 | 
|
307  | 
shows "([]::'x prm)\<bullet>x = x"  | 
|
308  | 
using a by (simp add: at_def)  | 
|
309  | 
||
310  | 
lemma at2:  | 
|
311  | 
fixes a ::"'x"  | 
|
312  | 
and b ::"'x"  | 
|
313  | 
and x ::"'x"  | 
|
314  | 
and pi::"'x prm"  | 
|
315  | 
  assumes a: "at TYPE('x)"
 | 
|
316  | 
shows "((a,b)#pi)\<bullet>x = swap (a,b) (pi\<bullet>x)"  | 
|
317  | 
using a by (simp only: at_def)  | 
|
318  | 
||
319  | 
lemma at3:  | 
|
320  | 
fixes a ::"'x"  | 
|
321  | 
and b ::"'x"  | 
|
322  | 
and c ::"'x"  | 
|
323  | 
  assumes a: "at TYPE('x)"
 | 
|
324  | 
shows "swap (a,b) c = (if a=c then b else (if b=c then a else c))"  | 
|
325  | 
using a by (simp only: at_def)  | 
|
326  | 
||
327  | 
(* rules to calculate simple premutations *)  | 
|
328  | 
lemmas at_calc = at2 at1 at3  | 
|
329  | 
||
330  | 
lemma at4:  | 
|
331  | 
  assumes a: "at TYPE('x)"
 | 
|
332  | 
shows "infinite (UNIV::'x set)"  | 
|
333  | 
using a by (simp add: at_def)  | 
|
334  | 
||
335  | 
lemma at_append:  | 
|
336  | 
fixes pi1 :: "'x prm"  | 
|
337  | 
and pi2 :: "'x prm"  | 
|
338  | 
and c :: "'x"  | 
|
339  | 
  assumes at: "at TYPE('x)" 
 | 
|
340  | 
shows "(pi1@pi2)\<bullet>c = pi1\<bullet>(pi2\<bullet>c)"  | 
|
341  | 
proof (induct pi1)  | 
|
342  | 
case Nil show ?case by (simp add: at1[OF at])  | 
|
343  | 
next  | 
|
344  | 
case (Cons x xs)  | 
|
| 
18053
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
345  | 
have "(xs@pi2)\<bullet>c = xs\<bullet>(pi2\<bullet>c)" by fact  | 
| 
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
346  | 
also have "(x#xs)@pi2 = x#(xs@pi2)" by simp  | 
| 
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
347  | 
ultimately show ?case by (cases "x", simp add: at2[OF at])  | 
| 17870 | 348  | 
qed  | 
349  | 
||
350  | 
lemma at_swap:  | 
|
351  | 
fixes a :: "'x"  | 
|
352  | 
and b :: "'x"  | 
|
353  | 
and c :: "'x"  | 
|
354  | 
  assumes at: "at TYPE('x)" 
 | 
|
355  | 
shows "swap (a,b) (swap (a,b) c) = c"  | 
|
356  | 
by (auto simp add: at3[OF at])  | 
|
357  | 
||
358  | 
lemma at_rev_pi:  | 
|
359  | 
fixes pi :: "'x prm"  | 
|
360  | 
and c :: "'x"  | 
|
361  | 
  assumes at: "at TYPE('x)"
 | 
|
362  | 
shows "(rev pi)\<bullet>(pi\<bullet>c) = c"  | 
|
363  | 
proof(induct pi)  | 
|
364  | 
case Nil show ?case by (simp add: at1[OF at])  | 
|
365  | 
next  | 
|
366  | 
case (Cons x xs) thus ?case  | 
|
367  | 
by (cases "x", simp add: at2[OF at] at_append[OF at] at1[OF at] at_swap[OF at])  | 
|
368  | 
qed  | 
|
369  | 
||
370  | 
lemma at_pi_rev:  | 
|
371  | 
fixes pi :: "'x prm"  | 
|
372  | 
and x :: "'x"  | 
|
373  | 
  assumes at: "at TYPE('x)"
 | 
|
374  | 
shows "pi\<bullet>((rev pi)\<bullet>x) = x"  | 
|
375  | 
by (rule at_rev_pi[OF at, of "rev pi" _,simplified])  | 
|
376  | 
||
377  | 
lemma at_bij1:  | 
|
378  | 
fixes pi :: "'x prm"  | 
|
379  | 
and x :: "'x"  | 
|
380  | 
and y :: "'x"  | 
|
381  | 
  assumes at: "at TYPE('x)"
 | 
|
382  | 
and a: "(pi\<bullet>x) = y"  | 
|
383  | 
shows "x=(rev pi)\<bullet>y"  | 
|
384  | 
proof -  | 
|
385  | 
from a have "y=(pi\<bullet>x)" by (rule sym)  | 
|
386  | 
thus ?thesis by (simp only: at_rev_pi[OF at])  | 
|
387  | 
qed  | 
|
388  | 
||
389  | 
lemma at_bij2:  | 
|
390  | 
fixes pi :: "'x prm"  | 
|
391  | 
and x :: "'x"  | 
|
392  | 
and y :: "'x"  | 
|
393  | 
  assumes at: "at TYPE('x)"
 | 
|
394  | 
and a: "((rev pi)\<bullet>x) = y"  | 
|
395  | 
shows "x=pi\<bullet>y"  | 
|
396  | 
proof -  | 
|
397  | 
from a have "y=((rev pi)\<bullet>x)" by (rule sym)  | 
|
398  | 
thus ?thesis by (simp only: at_pi_rev[OF at])  | 
|
399  | 
qed  | 
|
400  | 
||
401  | 
lemma at_bij:  | 
|
402  | 
fixes pi :: "'x prm"  | 
|
403  | 
and x :: "'x"  | 
|
404  | 
and y :: "'x"  | 
|
405  | 
  assumes at: "at TYPE('x)"
 | 
|
406  | 
shows "(pi\<bullet>x = pi\<bullet>y) = (x=y)"  | 
|
407  | 
proof  | 
|
408  | 
assume "pi\<bullet>x = pi\<bullet>y"  | 
|
409  | 
hence "x=(rev pi)\<bullet>(pi\<bullet>y)" by (rule at_bij1[OF at])  | 
|
410  | 
thus "x=y" by (simp only: at_rev_pi[OF at])  | 
|
411  | 
next  | 
|
412  | 
assume "x=y"  | 
|
413  | 
thus "pi\<bullet>x = pi\<bullet>y" by simp  | 
|
414  | 
qed  | 
|
415  | 
||
416  | 
lemma at_supp:  | 
|
417  | 
fixes x :: "'x"  | 
|
418  | 
  assumes at: "at TYPE('x)"
 | 
|
419  | 
  shows "supp x = {x}"
 | 
|
420  | 
proof (simp add: supp_def Collect_conj_eq Collect_imp_eq at_calc[OF at], auto)  | 
|
421  | 
  assume f: "finite {b::'x. b \<noteq> x}"
 | 
|
422  | 
  have a1: "{b::'x. b \<noteq> x} = UNIV-{x}" by force
 | 
|
423  | 
have a2: "infinite (UNIV::'x set)" by (rule at4[OF at])  | 
|
424  | 
from f a1 a2 show False by force  | 
|
425  | 
qed  | 
|
426  | 
||
427  | 
lemma at_fresh:  | 
|
428  | 
fixes a :: "'x"  | 
|
429  | 
and b :: "'x"  | 
|
430  | 
  assumes at: "at TYPE('x)"
 | 
|
431  | 
shows "(a\<sharp>b) = (a\<noteq>b)"  | 
|
432  | 
by (simp add: at_supp[OF at] fresh_def)  | 
|
433  | 
||
434  | 
lemma at_prm_fresh[rule_format]:  | 
|
435  | 
fixes c :: "'x"  | 
|
436  | 
and pi:: "'x prm"  | 
|
437  | 
  assumes at: "at TYPE('x)"
 | 
|
438  | 
shows "c\<sharp>pi \<longrightarrow> pi\<bullet>c = c"  | 
|
439  | 
apply(induct pi)  | 
|
440  | 
apply(simp add: at1[OF at])  | 
|
441  | 
apply(force simp add: fresh_list_cons at2[OF at] fresh_prod at_fresh[OF at] at3[OF at])  | 
|
442  | 
done  | 
|
443  | 
||
444  | 
lemma at_prm_rev_eq:  | 
|
445  | 
fixes pi1 :: "'x prm"  | 
|
446  | 
and pi2 :: "'x prm"  | 
|
447  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
448  | 
shows a: "((rev pi1) \<triangleq> (rev pi2)) = (pi1 \<triangleq> pi2)"  | 
| 17870 | 449  | 
proof (simp add: prm_eq_def, auto)  | 
450  | 
fix x  | 
|
451  | 
assume "\<forall>x::'x. (rev pi1)\<bullet>x = (rev pi2)\<bullet>x"  | 
|
452  | 
hence "(rev (pi1::'x prm))\<bullet>(pi2\<bullet>(x::'x)) = (rev (pi2::'x prm))\<bullet>(pi2\<bullet>x)" by simp  | 
|
453  | 
hence "(rev (pi1::'x prm))\<bullet>((pi2::'x prm)\<bullet>x) = (x::'x)" by (simp add: at_rev_pi[OF at])  | 
|
454  | 
hence "(pi2::'x prm)\<bullet>x = (pi1::'x prm)\<bullet>x" by (simp add: at_bij2[OF at])  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
455  | 
thus "pi1\<bullet>x = pi2\<bullet>x" by simp  | 
| 17870 | 456  | 
next  | 
457  | 
fix x  | 
|
458  | 
assume "\<forall>x::'x. pi1\<bullet>x = pi2\<bullet>x"  | 
|
459  | 
hence "(pi1::'x prm)\<bullet>((rev pi2)\<bullet>x) = (pi2::'x prm)\<bullet>((rev pi2)\<bullet>(x::'x))" by simp  | 
|
460  | 
hence "(pi1::'x prm)\<bullet>((rev pi2)\<bullet>(x::'x)) = x" by (simp add: at_pi_rev[OF at])  | 
|
461  | 
hence "(rev pi2)\<bullet>x = (rev pi1)\<bullet>(x::'x)" by (simp add: at_bij1[OF at])  | 
|
462  | 
thus "(rev pi1)\<bullet>x = (rev pi2)\<bullet>(x::'x)" by simp  | 
|
463  | 
qed  | 
|
464  | 
||
465  | 
lemma at_prm_rev_eq1:  | 
|
466  | 
fixes pi1 :: "'x prm"  | 
|
467  | 
and pi2 :: "'x prm"  | 
|
468  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
469  | 
shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1) \<triangleq> (rev pi2)"  | 
| 17870 | 470  | 
by (simp add: at_prm_rev_eq[OF at])  | 
471  | 
||
472  | 
lemma at_ds1:  | 
|
473  | 
fixes a :: "'x"  | 
|
474  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
475  | 
shows "[(a,a)] \<triangleq> []"  | 
| 17870 | 476  | 
by (force simp add: prm_eq_def at_calc[OF at])  | 
477  | 
||
478  | 
lemma at_ds2:  | 
|
479  | 
fixes pi :: "'x prm"  | 
|
480  | 
and a :: "'x"  | 
|
481  | 
and b :: "'x"  | 
|
482  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
483  | 
shows "(pi@[((rev pi)\<bullet>a,(rev pi)\<bullet>b)]) \<triangleq> ([(a,b)]@pi)"  | 
| 17870 | 484  | 
by (force simp add: prm_eq_def at_append[OF at] at_bij[OF at] at_pi_rev[OF at]  | 
485  | 
at_rev_pi[OF at] at_calc[OF at])  | 
|
486  | 
||
487  | 
lemma at_ds3:  | 
|
488  | 
fixes a :: "'x"  | 
|
489  | 
and b :: "'x"  | 
|
490  | 
and c :: "'x"  | 
|
491  | 
  assumes at: "at TYPE('x)"
 | 
|
492  | 
and a: "distinct [a,b,c]"  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
493  | 
shows "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]"  | 
| 17870 | 494  | 
using a by (force simp add: prm_eq_def at_calc[OF at])  | 
495  | 
||
496  | 
lemma at_ds4:  | 
|
497  | 
fixes a :: "'x"  | 
|
498  | 
and b :: "'x"  | 
|
499  | 
and pi :: "'x prm"  | 
|
500  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
501  | 
shows "(pi@[(a,(rev pi)\<bullet>b)]) \<triangleq> ([(pi\<bullet>a,b)]@pi)"  | 
| 17870 | 502  | 
by (force simp add: prm_eq_def at_append[OF at] at_calc[OF at] at_bij[OF at]  | 
503  | 
at_pi_rev[OF at] at_rev_pi[OF at])  | 
|
504  | 
||
505  | 
lemma at_ds5:  | 
|
506  | 
fixes a :: "'x"  | 
|
507  | 
and b :: "'x"  | 
|
508  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
509  | 
shows "[(a,b)] \<triangleq> [(b,a)]"  | 
| 17870 | 510  | 
by (force simp add: prm_eq_def at_calc[OF at])  | 
511  | 
||
512  | 
lemma at_ds6:  | 
|
513  | 
fixes a :: "'x"  | 
|
514  | 
and b :: "'x"  | 
|
515  | 
and c :: "'x"  | 
|
516  | 
  assumes at: "at TYPE('x)"
 | 
|
517  | 
and a: "distinct [a,b,c]"  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
518  | 
shows "[(a,c),(a,b)] \<triangleq> [(b,c),(a,c)]"  | 
| 17870 | 519  | 
using a by (force simp add: prm_eq_def at_calc[OF at])  | 
520  | 
||
521  | 
lemma at_ds7:  | 
|
522  | 
fixes pi :: "'x prm"  | 
|
523  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
524  | 
shows "((rev pi)@pi) \<triangleq> []"  | 
| 17870 | 525  | 
by (simp add: prm_eq_def at1[OF at] at_append[OF at] at_rev_pi[OF at])  | 
526  | 
||
527  | 
lemma at_ds8_aux:  | 
|
528  | 
fixes pi :: "'x prm"  | 
|
529  | 
and a :: "'x"  | 
|
530  | 
and b :: "'x"  | 
|
531  | 
and c :: "'x"  | 
|
532  | 
  assumes at: "at TYPE('x)"
 | 
|
533  | 
shows "pi\<bullet>(swap (a,b) c) = swap (pi\<bullet>a,pi\<bullet>b) (pi\<bullet>c)"  | 
|
534  | 
by (force simp add: at_calc[OF at] at_bij[OF at])  | 
|
535  | 
||
536  | 
lemma at_ds8:  | 
|
537  | 
fixes pi1 :: "'x prm"  | 
|
538  | 
and pi2 :: "'x prm"  | 
|
539  | 
and a :: "'x"  | 
|
540  | 
and b :: "'x"  | 
|
541  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
542  | 
shows "(pi1@pi2) \<triangleq> ((pi1\<bullet>pi2)@pi1)"  | 
| 17870 | 543  | 
apply(induct_tac pi2)  | 
544  | 
apply(simp add: prm_eq_def)  | 
|
545  | 
apply(auto simp add: prm_eq_def)  | 
|
546  | 
apply(simp add: at2[OF at])  | 
|
547  | 
apply(drule_tac x="aa" in spec)  | 
|
548  | 
apply(drule sym)  | 
|
549  | 
apply(simp)  | 
|
550  | 
apply(simp add: at_append[OF at])  | 
|
551  | 
apply(simp add: at2[OF at])  | 
|
552  | 
apply(simp add: at_ds8_aux[OF at])  | 
|
553  | 
done  | 
|
554  | 
||
555  | 
lemma at_ds9:  | 
|
556  | 
fixes pi1 :: "'x prm"  | 
|
557  | 
and pi2 :: "'x prm"  | 
|
558  | 
and a :: "'x"  | 
|
559  | 
and b :: "'x"  | 
|
560  | 
  assumes at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
561  | 
shows " ((rev pi2)@(rev pi1)) \<triangleq> ((rev pi1)@(rev (pi1\<bullet>pi2)))"  | 
| 17870 | 562  | 
apply(induct_tac pi2)  | 
563  | 
apply(simp add: prm_eq_def)  | 
|
564  | 
apply(auto simp add: prm_eq_def)  | 
|
565  | 
apply(simp add: at_append[OF at])  | 
|
566  | 
apply(simp add: at2[OF at] at1[OF at])  | 
|
567  | 
apply(drule_tac x="swap(pi1\<bullet>a,pi1\<bullet>b) aa" in spec)  | 
|
568  | 
apply(drule sym)  | 
|
569  | 
apply(simp)  | 
|
570  | 
apply(simp add: at_ds8_aux[OF at])  | 
|
571  | 
apply(simp add: at_rev_pi[OF at])  | 
|
572  | 
done  | 
|
573  | 
||
574  | 
--"there always exists an atom not being in a finite set"  | 
|
575  | 
lemma ex_in_inf:  | 
|
576  | 
fixes A::"'x set"  | 
|
577  | 
  assumes at: "at TYPE('x)"
 | 
|
578  | 
and fs: "finite A"  | 
|
579  | 
shows "\<exists>c::'x. c\<notin>A"  | 
|
580  | 
proof -  | 
|
581  | 
from fs at4[OF at] have "infinite ((UNIV::'x set) - A)"  | 
|
582  | 
by (simp add: Diff_infinite_finite)  | 
|
583  | 
  hence "((UNIV::'x set) - A) \<noteq> ({}::'x set)" by (force simp only:)
 | 
|
584  | 
hence "\<exists>c::'x. c\<in>((UNIV::'x set) - A)" by force  | 
|
585  | 
thus "\<exists>c::'x. c\<notin>A" by force  | 
|
586  | 
qed  | 
|
587  | 
||
588  | 
--"there always exists a fresh name for an object with finite support"  | 
|
589  | 
lemma at_exists_fresh:  | 
|
590  | 
fixes x :: "'a"  | 
|
591  | 
  assumes at: "at TYPE('x)"
 | 
|
592  | 
and fs: "finite ((supp x)::'x set)"  | 
|
593  | 
shows "\<exists>c::'x. c\<sharp>x"  | 
|
594  | 
by (simp add: fresh_def, rule ex_in_inf[OF at, OF fs])  | 
|
595  | 
||
596  | 
--"the at-props imply the pt-props"  | 
|
597  | 
lemma at_pt_inst:  | 
|
598  | 
  assumes at: "at TYPE('x)"
 | 
|
599  | 
  shows "pt TYPE('x) TYPE('x)"
 | 
|
600  | 
apply(auto simp only: pt_def)  | 
|
601  | 
apply(simp only: at1[OF at])  | 
|
602  | 
apply(simp only: at_append[OF at])  | 
|
| 
18053
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
603  | 
apply(simp only: prm_eq_def)  | 
| 17870 | 604  | 
done  | 
605  | 
||
606  | 
section {* finite support properties *}
 | 
|
607  | 
(*===================================*)  | 
|
608  | 
||
609  | 
lemma fs1:  | 
|
610  | 
fixes x :: "'a"  | 
|
611  | 
  assumes a: "fs TYPE('a) TYPE('x)"
 | 
|
612  | 
shows "finite ((supp x)::'x set)"  | 
|
613  | 
using a by (simp add: fs_def)  | 
|
614  | 
||
615  | 
lemma fs_at_inst:  | 
|
616  | 
fixes a :: "'x"  | 
|
617  | 
  assumes at: "at TYPE('x)"
 | 
|
618  | 
  shows "fs TYPE('x) TYPE('x)"
 | 
|
619  | 
apply(simp add: fs_def)  | 
|
620  | 
apply(simp add: at_supp[OF at])  | 
|
621  | 
done  | 
|
622  | 
||
623  | 
lemma fs_unit_inst:  | 
|
624  | 
  shows "fs TYPE(unit) TYPE('x)"
 | 
|
625  | 
apply(simp add: fs_def)  | 
|
626  | 
apply(simp add: supp_unit)  | 
|
627  | 
done  | 
|
628  | 
||
629  | 
lemma fs_prod_inst:  | 
|
630  | 
  assumes fsa: "fs TYPE('a) TYPE('x)"
 | 
|
631  | 
  and     fsb: "fs TYPE('b) TYPE('x)"
 | 
|
632  | 
  shows "fs TYPE('a\<times>'b) TYPE('x)"
 | 
|
633  | 
apply(unfold fs_def)  | 
|
634  | 
apply(auto simp add: supp_prod)  | 
|
635  | 
apply(rule fs1[OF fsa])  | 
|
636  | 
apply(rule fs1[OF fsb])  | 
|
637  | 
done  | 
|
638  | 
||
| 18600 | 639  | 
lemma fs_nprod_inst:  | 
640  | 
  assumes fsa: "fs TYPE('a) TYPE('x)"
 | 
|
641  | 
  and     fsb: "fs TYPE('b) TYPE('x)"
 | 
|
642  | 
  shows "fs TYPE(('a,'b) nprod) TYPE('x)"
 | 
|
643  | 
apply(unfold fs_def, rule allI)  | 
|
644  | 
apply(case_tac x)  | 
|
645  | 
apply(auto simp add: supp_nprod)  | 
|
646  | 
apply(rule fs1[OF fsa])  | 
|
647  | 
apply(rule fs1[OF fsb])  | 
|
648  | 
done  | 
|
649  | 
||
| 17870 | 650  | 
lemma fs_list_inst:  | 
651  | 
  assumes fs: "fs TYPE('a) TYPE('x)"
 | 
|
652  | 
  shows "fs TYPE('a list) TYPE('x)"
 | 
|
653  | 
apply(simp add: fs_def, rule allI)  | 
|
654  | 
apply(induct_tac x)  | 
|
655  | 
apply(simp add: supp_list_nil)  | 
|
656  | 
apply(simp add: supp_list_cons)  | 
|
657  | 
apply(rule fs1[OF fs])  | 
|
658  | 
done  | 
|
659  | 
||
| 18431 | 660  | 
lemma fs_option_inst:  | 
661  | 
  assumes fs: "fs TYPE('a) TYPE('x)"
 | 
|
662  | 
  shows "fs TYPE('a option) TYPE('x)"
 | 
|
| 17870 | 663  | 
apply(simp add: fs_def, rule allI)  | 
| 18431 | 664  | 
apply(case_tac x)  | 
665  | 
apply(simp add: supp_none)  | 
|
666  | 
apply(simp add: supp_some)  | 
|
667  | 
apply(rule fs1[OF fs])  | 
|
| 17870 | 668  | 
done  | 
669  | 
||
670  | 
section {* Lemmas about the permutation properties *}
 | 
|
671  | 
(*=================================================*)  | 
|
672  | 
||
673  | 
lemma pt1:  | 
|
674  | 
fixes x::"'a"  | 
|
675  | 
  assumes a: "pt TYPE('a) TYPE('x)"
 | 
|
676  | 
shows "([]::'x prm)\<bullet>x = x"  | 
|
677  | 
using a by (simp add: pt_def)  | 
|
678  | 
||
679  | 
lemma pt2:  | 
|
680  | 
fixes pi1::"'x prm"  | 
|
681  | 
and pi2::"'x prm"  | 
|
682  | 
and x ::"'a"  | 
|
683  | 
  assumes a: "pt TYPE('a) TYPE('x)"
 | 
|
684  | 
shows "(pi1@pi2)\<bullet>x = pi1\<bullet>(pi2\<bullet>x)"  | 
|
685  | 
using a by (simp add: pt_def)  | 
|
686  | 
||
687  | 
lemma pt3:  | 
|
688  | 
fixes pi1::"'x prm"  | 
|
689  | 
and pi2::"'x prm"  | 
|
690  | 
and x ::"'a"  | 
|
691  | 
  assumes a: "pt TYPE('a) TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
692  | 
shows "pi1 \<triangleq> pi2 \<Longrightarrow> pi1\<bullet>x = pi2\<bullet>x"  | 
| 17870 | 693  | 
using a by (simp add: pt_def)  | 
694  | 
||
695  | 
lemma pt3_rev:  | 
|
696  | 
fixes pi1::"'x prm"  | 
|
697  | 
and pi2::"'x prm"  | 
|
698  | 
and x ::"'a"  | 
|
699  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
700  | 
  and     at: "at TYPE('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
701  | 
shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1)\<bullet>x = (rev pi2)\<bullet>x"  | 
| 17870 | 702  | 
by (rule pt3[OF pt], simp add: at_prm_rev_eq[OF at])  | 
703  | 
||
704  | 
section {* composition properties *}
 | 
|
705  | 
(* ============================== *)  | 
|
706  | 
lemma cp1:  | 
|
707  | 
fixes pi1::"'x prm"  | 
|
708  | 
and pi2::"'y prm"  | 
|
709  | 
and x ::"'a"  | 
|
710  | 
  assumes cp: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
711  | 
shows "pi1\<bullet>(pi2\<bullet>x) = (pi1\<bullet>pi2)\<bullet>(pi1\<bullet>x)"  | 
|
712  | 
using cp by (simp add: cp_def)  | 
|
713  | 
||
714  | 
lemma cp_pt_inst:  | 
|
715  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
716  | 
  and     at: "at TYPE('x)"
 | 
|
717  | 
  shows "cp TYPE('a) TYPE('x) TYPE('x)"
 | 
|
718  | 
apply(auto simp add: cp_def pt2[OF pt,symmetric])  | 
|
719  | 
apply(rule pt3[OF pt])  | 
|
720  | 
apply(rule at_ds8[OF at])  | 
|
721  | 
done  | 
|
722  | 
||
723  | 
section {* permutation type instances *}
 | 
|
724  | 
(* ===================================*)  | 
|
725  | 
||
726  | 
lemma pt_set_inst:  | 
|
727  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
728  | 
  shows  "pt TYPE('a set) TYPE('x)"
 | 
|
729  | 
apply(simp add: pt_def)  | 
|
730  | 
apply(simp_all add: perm_set_def)  | 
|
731  | 
apply(simp add: pt1[OF pt])  | 
|
732  | 
apply(force simp add: pt2[OF pt] pt3[OF pt])  | 
|
733  | 
done  | 
|
734  | 
||
735  | 
lemma pt_list_nil:  | 
|
736  | 
fixes xs :: "'a list"  | 
|
737  | 
  assumes pt: "pt TYPE('a) TYPE ('x)"
 | 
|
738  | 
shows "([]::'x prm)\<bullet>xs = xs"  | 
|
739  | 
apply(induct_tac xs)  | 
|
740  | 
apply(simp_all add: pt1[OF pt])  | 
|
741  | 
done  | 
|
742  | 
||
743  | 
lemma pt_list_append:  | 
|
744  | 
fixes pi1 :: "'x prm"  | 
|
745  | 
and pi2 :: "'x prm"  | 
|
746  | 
and xs :: "'a list"  | 
|
747  | 
  assumes pt: "pt TYPE('a) TYPE ('x)"
 | 
|
748  | 
shows "(pi1@pi2)\<bullet>xs = pi1\<bullet>(pi2\<bullet>xs)"  | 
|
749  | 
apply(induct_tac xs)  | 
|
750  | 
apply(simp_all add: pt2[OF pt])  | 
|
751  | 
done  | 
|
752  | 
||
753  | 
lemma pt_list_prm_eq:  | 
|
754  | 
fixes pi1 :: "'x prm"  | 
|
755  | 
and pi2 :: "'x prm"  | 
|
756  | 
and xs :: "'a list"  | 
|
757  | 
  assumes pt: "pt TYPE('a) TYPE ('x)"
 | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
758  | 
shows "pi1 \<triangleq> pi2 \<Longrightarrow> pi1\<bullet>xs = pi2\<bullet>xs"  | 
| 17870 | 759  | 
apply(induct_tac xs)  | 
760  | 
apply(simp_all add: prm_eq_def pt3[OF pt])  | 
|
761  | 
done  | 
|
762  | 
||
763  | 
lemma pt_list_inst:  | 
|
764  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
765  | 
  shows  "pt TYPE('a list) TYPE('x)"
 | 
|
766  | 
apply(auto simp only: pt_def)  | 
|
767  | 
apply(rule pt_list_nil[OF pt])  | 
|
768  | 
apply(rule pt_list_append[OF pt])  | 
|
769  | 
apply(rule pt_list_prm_eq[OF pt],assumption)  | 
|
770  | 
done  | 
|
771  | 
||
772  | 
lemma pt_unit_inst:  | 
|
773  | 
  shows  "pt TYPE(unit) TYPE('x)"
 | 
|
774  | 
by (simp add: pt_def)  | 
|
775  | 
||
776  | 
lemma pt_prod_inst:  | 
|
777  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
778  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
779  | 
  shows  "pt TYPE('a \<times> 'b) TYPE('x)"
 | 
|
780  | 
apply(auto simp add: pt_def)  | 
|
781  | 
apply(rule pt1[OF pta])  | 
|
782  | 
apply(rule pt1[OF ptb])  | 
|
783  | 
apply(rule pt2[OF pta])  | 
|
784  | 
apply(rule pt2[OF ptb])  | 
|
785  | 
apply(rule pt3[OF pta],assumption)  | 
|
786  | 
apply(rule pt3[OF ptb],assumption)  | 
|
787  | 
done  | 
|
788  | 
||
| 18600 | 789  | 
lemma pt_nprod_inst:  | 
790  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
791  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
792  | 
  shows  "pt TYPE(('a,'b) nprod) TYPE('x)"
 | 
|
793  | 
apply(auto simp add: pt_def)  | 
|
794  | 
apply(case_tac x)  | 
|
795  | 
apply(simp add: pt1[OF pta] pt1[OF ptb])  | 
|
796  | 
apply(case_tac x)  | 
|
797  | 
apply(simp add: pt2[OF pta] pt2[OF ptb])  | 
|
798  | 
apply(case_tac x)  | 
|
799  | 
apply(simp add: pt3[OF pta] pt3[OF ptb])  | 
|
800  | 
done  | 
|
801  | 
||
| 17870 | 802  | 
lemma pt_fun_inst:  | 
803  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
804  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
805  | 
  and     at:  "at TYPE('x)"
 | 
|
806  | 
  shows  "pt TYPE('a\<Rightarrow>'b) TYPE('x)"
 | 
|
807  | 
apply(auto simp only: pt_def)  | 
|
808  | 
apply(simp_all add: perm_fun_def)  | 
|
809  | 
apply(simp add: pt1[OF pta] pt1[OF ptb])  | 
|
810  | 
apply(simp add: pt2[OF pta] pt2[OF ptb])  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
811  | 
apply(subgoal_tac "(rev pi1) \<triangleq> (rev pi2)")(*A*)  | 
| 17870 | 812  | 
apply(simp add: pt3[OF pta] pt3[OF ptb])  | 
813  | 
(*A*)  | 
|
814  | 
apply(simp add: at_prm_rev_eq[OF at])  | 
|
815  | 
done  | 
|
816  | 
||
817  | 
lemma pt_option_inst:  | 
|
818  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
819  | 
  shows  "pt TYPE('a option) TYPE('x)"
 | 
|
820  | 
apply(auto simp only: pt_def)  | 
|
821  | 
apply(case_tac "x")  | 
|
822  | 
apply(simp_all add: pt1[OF pta])  | 
|
823  | 
apply(case_tac "x")  | 
|
824  | 
apply(simp_all add: pt2[OF pta])  | 
|
825  | 
apply(case_tac "x")  | 
|
826  | 
apply(simp_all add: pt3[OF pta])  | 
|
827  | 
done  | 
|
828  | 
||
829  | 
lemma pt_noption_inst:  | 
|
830  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
831  | 
  shows  "pt TYPE('a noption) TYPE('x)"
 | 
| 17870 | 832  | 
apply(auto simp only: pt_def)  | 
833  | 
apply(case_tac "x")  | 
|
834  | 
apply(simp_all add: pt1[OF pta])  | 
|
835  | 
apply(case_tac "x")  | 
|
836  | 
apply(simp_all add: pt2[OF pta])  | 
|
837  | 
apply(case_tac "x")  | 
|
838  | 
apply(simp_all add: pt3[OF pta])  | 
|
839  | 
done  | 
|
840  | 
||
841  | 
section {* further lemmas for permutation types *}
 | 
|
842  | 
(*==============================================*)  | 
|
843  | 
||
844  | 
lemma pt_rev_pi:  | 
|
845  | 
fixes pi :: "'x prm"  | 
|
846  | 
and x :: "'a"  | 
|
847  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
848  | 
  and     at: "at TYPE('x)"
 | 
|
849  | 
shows "(rev pi)\<bullet>(pi\<bullet>x) = x"  | 
|
850  | 
proof -  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
851  | 
have "((rev pi)@pi) \<triangleq> ([]::'x prm)" by (simp add: at_ds7[OF at])  | 
| 17870 | 852  | 
hence "((rev pi)@pi)\<bullet>(x::'a) = ([]::'x prm)\<bullet>x" by (simp add: pt3[OF pt])  | 
853  | 
thus ?thesis by (simp add: pt1[OF pt] pt2[OF pt])  | 
|
854  | 
qed  | 
|
855  | 
||
856  | 
lemma pt_pi_rev:  | 
|
857  | 
fixes pi :: "'x prm"  | 
|
858  | 
and x :: "'a"  | 
|
859  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
860  | 
  and     at: "at TYPE('x)"
 | 
|
861  | 
shows "pi\<bullet>((rev pi)\<bullet>x) = x"  | 
|
862  | 
by (simp add: pt_rev_pi[OF pt, OF at,of "rev pi" "x",simplified])  | 
|
863  | 
||
864  | 
lemma pt_bij1:  | 
|
865  | 
fixes pi :: "'x prm"  | 
|
866  | 
and x :: "'a"  | 
|
867  | 
and y :: "'a"  | 
|
868  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
869  | 
  and     at: "at TYPE('x)"
 | 
|
870  | 
and a: "(pi\<bullet>x) = y"  | 
|
871  | 
shows "x=(rev pi)\<bullet>y"  | 
|
872  | 
proof -  | 
|
873  | 
from a have "y=(pi\<bullet>x)" by (rule sym)  | 
|
874  | 
thus ?thesis by (simp only: pt_rev_pi[OF pt, OF at])  | 
|
875  | 
qed  | 
|
876  | 
||
877  | 
lemma pt_bij2:  | 
|
878  | 
fixes pi :: "'x prm"  | 
|
879  | 
and x :: "'a"  | 
|
880  | 
and y :: "'a"  | 
|
881  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
882  | 
  and     at: "at TYPE('x)"
 | 
|
883  | 
and a: "x = (rev pi)\<bullet>y"  | 
|
884  | 
shows "(pi\<bullet>x)=y"  | 
|
885  | 
using a by (simp add: pt_pi_rev[OF pt, OF at])  | 
|
886  | 
||
887  | 
lemma pt_bij:  | 
|
888  | 
fixes pi :: "'x prm"  | 
|
889  | 
and x :: "'a"  | 
|
890  | 
and y :: "'a"  | 
|
891  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
892  | 
  and     at: "at TYPE('x)"
 | 
|
893  | 
shows "(pi\<bullet>x = pi\<bullet>y) = (x=y)"  | 
|
894  | 
proof  | 
|
895  | 
assume "pi\<bullet>x = pi\<bullet>y"  | 
|
896  | 
hence "x=(rev pi)\<bullet>(pi\<bullet>y)" by (rule pt_bij1[OF pt, OF at])  | 
|
897  | 
thus "x=y" by (simp only: pt_rev_pi[OF pt, OF at])  | 
|
898  | 
next  | 
|
899  | 
assume "x=y"  | 
|
900  | 
thus "pi\<bullet>x = pi\<bullet>y" by simp  | 
|
901  | 
qed  | 
|
902  | 
||
903  | 
lemma pt_bij3:  | 
|
904  | 
fixes pi :: "'x prm"  | 
|
905  | 
and x :: "'a"  | 
|
906  | 
and y :: "'a"  | 
|
907  | 
assumes a: "x=y"  | 
|
908  | 
shows "(pi\<bullet>x = pi\<bullet>y)"  | 
|
909  | 
using a by simp  | 
|
910  | 
||
911  | 
lemma pt_bij4:  | 
|
912  | 
fixes pi :: "'x prm"  | 
|
913  | 
and x :: "'a"  | 
|
914  | 
and y :: "'a"  | 
|
915  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
916  | 
  and     at: "at TYPE('x)"
 | 
|
917  | 
and a: "pi\<bullet>x = pi\<bullet>y"  | 
|
918  | 
shows "x = y"  | 
|
919  | 
using a by (simp add: pt_bij[OF pt, OF at])  | 
|
920  | 
||
921  | 
lemma pt_swap_bij:  | 
|
922  | 
fixes a :: "'x"  | 
|
923  | 
and b :: "'x"  | 
|
924  | 
and x :: "'a"  | 
|
925  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
926  | 
  and     at: "at TYPE('x)"
 | 
|
927  | 
shows "[(a,b)]\<bullet>([(a,b)]\<bullet>x) = x"  | 
|
928  | 
by (rule pt_bij2[OF pt, OF at], simp)  | 
|
929  | 
||
930  | 
lemma pt_set_bij1:  | 
|
931  | 
fixes pi :: "'x prm"  | 
|
932  | 
and x :: "'a"  | 
|
933  | 
and X :: "'a set"  | 
|
934  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
935  | 
  and     at: "at TYPE('x)"
 | 
|
936  | 
shows "((pi\<bullet>x)\<in>X) = (x\<in>((rev pi)\<bullet>X))"  | 
|
937  | 
by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])  | 
|
938  | 
||
939  | 
lemma pt_set_bij1a:  | 
|
940  | 
fixes pi :: "'x prm"  | 
|
941  | 
and x :: "'a"  | 
|
942  | 
and X :: "'a set"  | 
|
943  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
944  | 
  and     at: "at TYPE('x)"
 | 
|
945  | 
shows "(x\<in>(pi\<bullet>X)) = (((rev pi)\<bullet>x)\<in>X)"  | 
|
946  | 
by (force simp add: perm_set_def pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])  | 
|
947  | 
||
948  | 
lemma pt_set_bij:  | 
|
949  | 
fixes pi :: "'x prm"  | 
|
950  | 
and x :: "'a"  | 
|
951  | 
and X :: "'a set"  | 
|
952  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
953  | 
  and     at: "at TYPE('x)"
 | 
|
954  | 
shows "((pi\<bullet>x)\<in>(pi\<bullet>X)) = (x\<in>X)"  | 
|
| 
18053
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
955  | 
by (simp add: perm_set_def pt_bij[OF pt, OF at])  | 
| 17870 | 956  | 
|
957  | 
lemma pt_set_bij2:  | 
|
958  | 
fixes pi :: "'x prm"  | 
|
959  | 
and x :: "'a"  | 
|
960  | 
and X :: "'a set"  | 
|
961  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
962  | 
  and     at: "at TYPE('x)"
 | 
|
963  | 
and a: "x\<in>X"  | 
|
964  | 
shows "(pi\<bullet>x)\<in>(pi\<bullet>X)"  | 
|
965  | 
using a by (simp add: pt_set_bij[OF pt, OF at])  | 
|
966  | 
||
| 18264 | 967  | 
lemma pt_set_bij2a:  | 
968  | 
fixes pi :: "'x prm"  | 
|
969  | 
and x :: "'a"  | 
|
970  | 
and X :: "'a set"  | 
|
971  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
972  | 
  and     at: "at TYPE('x)"
 | 
|
973  | 
and a: "x\<in>((rev pi)\<bullet>X)"  | 
|
974  | 
shows "(pi\<bullet>x)\<in>X"  | 
|
975  | 
using a by (simp add: pt_set_bij1[OF pt, OF at])  | 
|
976  | 
||
| 17870 | 977  | 
lemma pt_set_bij3:  | 
978  | 
fixes pi :: "'x prm"  | 
|
979  | 
and x :: "'a"  | 
|
980  | 
and X :: "'a set"  | 
|
981  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
982  | 
  and     at: "at TYPE('x)"
 | 
|
983  | 
shows "pi\<bullet>(x\<in>X) = (x\<in>X)"  | 
|
984  | 
apply(case_tac "x\<in>X = True")  | 
|
985  | 
apply(auto)  | 
|
986  | 
done  | 
|
987  | 
||
| 
18159
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
988  | 
lemma pt_subseteq_eqvt:  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
989  | 
fixes pi :: "'x prm"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
990  | 
and Y :: "'a set"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
991  | 
and X :: "'a set"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
992  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
993  | 
  and     at: "at TYPE('x)"
 | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
994  | 
shows "((pi\<bullet>X)\<subseteq>(pi\<bullet>Y)) = (X\<subseteq>Y)"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
995  | 
proof (auto)  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
996  | 
fix x::"'a"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
997  | 
assume a: "(pi\<bullet>X)\<subseteq>(pi\<bullet>Y)"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
998  | 
and "x\<in>X"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
999  | 
hence "(pi\<bullet>x)\<in>(pi\<bullet>X)" by (simp add: pt_set_bij[OF pt, OF at])  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1000  | 
with a have "(pi\<bullet>x)\<in>(pi\<bullet>Y)" by force  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1001  | 
thus "x\<in>Y" by (simp add: pt_set_bij[OF pt, OF at])  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1002  | 
next  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1003  | 
fix x::"'a"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1004  | 
assume a: "X\<subseteq>Y"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1005  | 
and "x\<in>(pi\<bullet>X)"  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1006  | 
thus "x\<in>(pi\<bullet>Y)" by (force simp add: pt_set_bij1a[OF pt, OF at])  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1007  | 
qed  | 
| 
 
08282ca0402e
added a few equivariance lemmas (they need to be automated
 
urbanc 
parents: 
18068 
diff
changeset
 | 
1008  | 
|
| 17870 | 1009  | 
-- "some helper lemmas for the pt_perm_supp_ineq lemma"  | 
1010  | 
lemma Collect_permI:  | 
|
1011  | 
fixes pi :: "'x prm"  | 
|
1012  | 
and x :: "'a"  | 
|
1013  | 
assumes a: "\<forall>x. (P1 x = P2 x)"  | 
|
1014  | 
  shows "{pi\<bullet>x| x. P1 x} = {pi\<bullet>x| x. P2 x}"
 | 
|
1015  | 
using a by force  | 
|
1016  | 
||
1017  | 
lemma Infinite_cong:  | 
|
1018  | 
assumes a: "X = Y"  | 
|
1019  | 
shows "infinite X = infinite Y"  | 
|
1020  | 
using a by (simp)  | 
|
1021  | 
||
1022  | 
lemma pt_set_eq_ineq:  | 
|
1023  | 
fixes pi :: "'y prm"  | 
|
1024  | 
  assumes pt: "pt TYPE('x) TYPE('y)"
 | 
|
1025  | 
  and     at: "at TYPE('y)"
 | 
|
1026  | 
  shows "{pi\<bullet>x| x::'x. P x} = {x::'x. P ((rev pi)\<bullet>x)}"
 | 
|
1027  | 
by (force simp only: pt_rev_pi[OF pt, OF at] pt_pi_rev[OF pt, OF at])  | 
|
1028  | 
||
1029  | 
lemma pt_inject_on_ineq:  | 
|
1030  | 
fixes X :: "'y set"  | 
|
1031  | 
and pi :: "'x prm"  | 
|
1032  | 
  assumes pt: "pt TYPE('y) TYPE('x)"
 | 
|
1033  | 
  and     at: "at TYPE('x)"
 | 
|
1034  | 
shows "inj_on (perm pi) X"  | 
|
1035  | 
proof (unfold inj_on_def, intro strip)  | 
|
1036  | 
fix x::"'y" and y::"'y"  | 
|
1037  | 
assume "pi\<bullet>x = pi\<bullet>y"  | 
|
1038  | 
thus "x=y" by (simp add: pt_bij[OF pt, OF at])  | 
|
1039  | 
qed  | 
|
1040  | 
||
1041  | 
lemma pt_set_finite_ineq:  | 
|
1042  | 
fixes X :: "'x set"  | 
|
1043  | 
and pi :: "'y prm"  | 
|
1044  | 
  assumes pt: "pt TYPE('x) TYPE('y)"
 | 
|
1045  | 
  and     at: "at TYPE('y)"
 | 
|
1046  | 
shows "finite (pi\<bullet>X) = finite X"  | 
|
1047  | 
proof -  | 
|
1048  | 
have image: "(pi\<bullet>X) = (perm pi ` X)" by (force simp only: perm_set_def)  | 
|
1049  | 
show ?thesis  | 
|
1050  | 
proof (rule iffI)  | 
|
1051  | 
assume "finite (pi\<bullet>X)"  | 
|
1052  | 
hence "finite (perm pi ` X)" using image by (simp)  | 
|
1053  | 
thus "finite X" using pt_inject_on_ineq[OF pt, OF at] by (rule finite_imageD)  | 
|
1054  | 
next  | 
|
1055  | 
assume "finite X"  | 
|
1056  | 
hence "finite (perm pi ` X)" by (rule finite_imageI)  | 
|
1057  | 
thus "finite (pi\<bullet>X)" using image by (simp)  | 
|
1058  | 
qed  | 
|
1059  | 
qed  | 
|
1060  | 
||
1061  | 
lemma pt_set_infinite_ineq:  | 
|
1062  | 
fixes X :: "'x set"  | 
|
1063  | 
and pi :: "'y prm"  | 
|
1064  | 
  assumes pt: "pt TYPE('x) TYPE('y)"
 | 
|
1065  | 
  and     at: "at TYPE('y)"
 | 
|
1066  | 
shows "infinite (pi\<bullet>X) = infinite X"  | 
|
1067  | 
using pt at by (simp add: pt_set_finite_ineq)  | 
|
1068  | 
||
1069  | 
lemma pt_perm_supp_ineq:  | 
|
1070  | 
fixes pi :: "'x prm"  | 
|
1071  | 
and x :: "'a"  | 
|
1072  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1073  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
1074  | 
  and     at:  "at TYPE('x)"
 | 
|
1075  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
1076  | 
shows "(pi\<bullet>((supp x)::'y set)) = supp (pi\<bullet>x)" (is "?LHS = ?RHS")  | 
|
1077  | 
proof -  | 
|
1078  | 
  have "?LHS = {pi\<bullet>a | a. infinite {b. [(a,b)]\<bullet>x \<noteq> x}}" by (simp add: supp_def perm_set_def)
 | 
|
1079  | 
  also have "\<dots> = {pi\<bullet>a | a. infinite {pi\<bullet>b | b. [(a,b)]\<bullet>x \<noteq> x}}" 
 | 
|
1080  | 
proof (rule Collect_permI, rule allI, rule iffI)  | 
|
1081  | 
fix a  | 
|
1082  | 
    assume "infinite {b::'y. [(a,b)]\<bullet>x  \<noteq> x}"
 | 
|
1083  | 
    hence "infinite (pi\<bullet>{b::'y. [(a,b)]\<bullet>x \<noteq> x})" by (simp add: pt_set_infinite_ineq[OF ptb, OF at])
 | 
|
1084  | 
    thus "infinite {pi\<bullet>b |b::'y. [(a,b)]\<bullet>x  \<noteq> x}" by (simp add: perm_set_def)
 | 
|
1085  | 
next  | 
|
1086  | 
fix a  | 
|
1087  | 
    assume "infinite {pi\<bullet>b |b::'y. [(a,b)]\<bullet>x \<noteq> x}"
 | 
|
1088  | 
    hence "infinite (pi\<bullet>{b::'y. [(a,b)]\<bullet>x \<noteq> x})" by (simp add: perm_set_def)
 | 
|
1089  | 
    thus "infinite {b::'y. [(a,b)]\<bullet>x  \<noteq> x}" 
 | 
|
1090  | 
by (simp add: pt_set_infinite_ineq[OF ptb, OF at])  | 
|
1091  | 
qed  | 
|
1092  | 
  also have "\<dots> = {a. infinite {b::'y. [((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x \<noteq> x}}" 
 | 
|
1093  | 
by (simp add: pt_set_eq_ineq[OF ptb, OF at])  | 
|
1094  | 
  also have "\<dots> = {a. infinite {b. pi\<bullet>([((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x) \<noteq> (pi\<bullet>x)}}"
 | 
|
1095  | 
by (simp add: pt_bij[OF pta, OF at])  | 
|
1096  | 
  also have "\<dots> = {a. infinite {b. [(a,b)]\<bullet>(pi\<bullet>x) \<noteq> (pi\<bullet>x)}}"
 | 
|
1097  | 
proof (rule Collect_cong, rule Infinite_cong, rule Collect_cong)  | 
|
1098  | 
fix a::"'y" and b::"'y"  | 
|
1099  | 
have "pi\<bullet>(([((rev pi)\<bullet>a,(rev pi)\<bullet>b)])\<bullet>x) = [(a,b)]\<bullet>(pi\<bullet>x)"  | 
|
1100  | 
by (simp add: cp1[OF cp] pt_pi_rev[OF ptb, OF at])  | 
|
1101  | 
thus "(pi\<bullet>([((rev pi)\<bullet>a,(rev pi)\<bullet>b)]\<bullet>x) \<noteq> pi\<bullet>x) = ([(a,b)]\<bullet>(pi\<bullet>x) \<noteq> pi\<bullet>x)" by simp  | 
|
1102  | 
qed  | 
|
1103  | 
finally show "?LHS = ?RHS" by (simp add: supp_def)  | 
|
1104  | 
qed  | 
|
1105  | 
||
1106  | 
lemma pt_perm_supp:  | 
|
1107  | 
fixes pi :: "'x prm"  | 
|
1108  | 
and x :: "'a"  | 
|
1109  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1110  | 
  and     at: "at TYPE('x)"
 | 
|
1111  | 
shows "(pi\<bullet>((supp x)::'x set)) = supp (pi\<bullet>x)"  | 
|
1112  | 
apply(rule pt_perm_supp_ineq)  | 
|
1113  | 
apply(rule pt)  | 
|
1114  | 
apply(rule at_pt_inst)  | 
|
1115  | 
apply(rule at)+  | 
|
1116  | 
apply(rule cp_pt_inst)  | 
|
1117  | 
apply(rule pt)  | 
|
1118  | 
apply(rule at)  | 
|
1119  | 
done  | 
|
1120  | 
||
1121  | 
lemma pt_supp_finite_pi:  | 
|
1122  | 
fixes pi :: "'x prm"  | 
|
1123  | 
and x :: "'a"  | 
|
1124  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1125  | 
  and     at: "at TYPE('x)"
 | 
|
1126  | 
and f: "finite ((supp x)::'x set)"  | 
|
1127  | 
shows "finite ((supp (pi\<bullet>x))::'x set)"  | 
|
1128  | 
apply(simp add: pt_perm_supp[OF pt, OF at, symmetric])  | 
|
1129  | 
apply(simp add: pt_set_finite_ineq[OF at_pt_inst[OF at], OF at])  | 
|
1130  | 
apply(rule f)  | 
|
1131  | 
done  | 
|
1132  | 
||
1133  | 
lemma pt_fresh_left_ineq:  | 
|
1134  | 
fixes pi :: "'x prm"  | 
|
1135  | 
and x :: "'a"  | 
|
1136  | 
and a :: "'y"  | 
|
1137  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1138  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
1139  | 
  and     at:  "at TYPE('x)"
 | 
|
1140  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
1141  | 
shows "a\<sharp>(pi\<bullet>x) = ((rev pi)\<bullet>a)\<sharp>x"  | 
|
1142  | 
apply(simp add: fresh_def)  | 
|
1143  | 
apply(simp add: pt_set_bij1[OF ptb, OF at])  | 
|
1144  | 
apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp])  | 
|
1145  | 
done  | 
|
1146  | 
||
1147  | 
lemma pt_fresh_right_ineq:  | 
|
1148  | 
fixes pi :: "'x prm"  | 
|
1149  | 
and x :: "'a"  | 
|
1150  | 
and a :: "'y"  | 
|
1151  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1152  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
1153  | 
  and     at:  "at TYPE('x)"
 | 
|
1154  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
1155  | 
shows "(pi\<bullet>a)\<sharp>x = a\<sharp>((rev pi)\<bullet>x)"  | 
|
1156  | 
apply(simp add: fresh_def)  | 
|
1157  | 
apply(simp add: pt_set_bij1[OF ptb, OF at])  | 
|
1158  | 
apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp])  | 
|
1159  | 
done  | 
|
1160  | 
||
1161  | 
lemma pt_fresh_bij_ineq:  | 
|
1162  | 
fixes pi :: "'x prm"  | 
|
1163  | 
and x :: "'a"  | 
|
1164  | 
and a :: "'y"  | 
|
1165  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1166  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
1167  | 
  and     at:  "at TYPE('x)"
 | 
|
1168  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
1169  | 
shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x) = a\<sharp>x"  | 
|
1170  | 
apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])  | 
|
1171  | 
apply(simp add: pt_rev_pi[OF ptb, OF at])  | 
|
1172  | 
done  | 
|
1173  | 
||
1174  | 
lemma pt_fresh_left:  | 
|
1175  | 
fixes pi :: "'x prm"  | 
|
1176  | 
and x :: "'a"  | 
|
1177  | 
and a :: "'x"  | 
|
1178  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1179  | 
  and     at: "at TYPE('x)"
 | 
|
1180  | 
shows "a\<sharp>(pi\<bullet>x) = ((rev pi)\<bullet>a)\<sharp>x"  | 
|
1181  | 
apply(rule pt_fresh_left_ineq)  | 
|
1182  | 
apply(rule pt)  | 
|
1183  | 
apply(rule at_pt_inst)  | 
|
1184  | 
apply(rule at)+  | 
|
1185  | 
apply(rule cp_pt_inst)  | 
|
1186  | 
apply(rule pt)  | 
|
1187  | 
apply(rule at)  | 
|
1188  | 
done  | 
|
1189  | 
||
1190  | 
lemma pt_fresh_right:  | 
|
1191  | 
fixes pi :: "'x prm"  | 
|
1192  | 
and x :: "'a"  | 
|
1193  | 
and a :: "'x"  | 
|
1194  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1195  | 
  and     at: "at TYPE('x)"
 | 
|
1196  | 
shows "(pi\<bullet>a)\<sharp>x = a\<sharp>((rev pi)\<bullet>x)"  | 
|
1197  | 
apply(rule pt_fresh_right_ineq)  | 
|
1198  | 
apply(rule pt)  | 
|
1199  | 
apply(rule at_pt_inst)  | 
|
1200  | 
apply(rule at)+  | 
|
1201  | 
apply(rule cp_pt_inst)  | 
|
1202  | 
apply(rule pt)  | 
|
1203  | 
apply(rule at)  | 
|
1204  | 
done  | 
|
1205  | 
||
1206  | 
lemma pt_fresh_bij:  | 
|
1207  | 
fixes pi :: "'x prm"  | 
|
1208  | 
and x :: "'a"  | 
|
1209  | 
and a :: "'x"  | 
|
1210  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1211  | 
  and     at: "at TYPE('x)"
 | 
|
1212  | 
shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x) = a\<sharp>x"  | 
|
1213  | 
apply(rule pt_fresh_bij_ineq)  | 
|
1214  | 
apply(rule pt)  | 
|
1215  | 
apply(rule at_pt_inst)  | 
|
1216  | 
apply(rule at)+  | 
|
1217  | 
apply(rule cp_pt_inst)  | 
|
1218  | 
apply(rule pt)  | 
|
1219  | 
apply(rule at)  | 
|
1220  | 
done  | 
|
1221  | 
||
1222  | 
lemma pt_fresh_bij1:  | 
|
1223  | 
fixes pi :: "'x prm"  | 
|
1224  | 
and x :: "'a"  | 
|
1225  | 
and a :: "'x"  | 
|
1226  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1227  | 
  and     at: "at TYPE('x)"
 | 
|
1228  | 
and a: "a\<sharp>x"  | 
|
1229  | 
shows "(pi\<bullet>a)\<sharp>(pi\<bullet>x)"  | 
|
1230  | 
using a by (simp add: pt_fresh_bij[OF pt, OF at])  | 
|
1231  | 
||
1232  | 
lemma pt_perm_fresh1:  | 
|
1233  | 
fixes a :: "'x"  | 
|
1234  | 
and b :: "'x"  | 
|
1235  | 
and x :: "'a"  | 
|
1236  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1237  | 
  and     at: "at TYPE ('x)"
 | 
|
1238  | 
and a1: "\<not>(a\<sharp>x)"  | 
|
1239  | 
and a2: "b\<sharp>x"  | 
|
1240  | 
shows "[(a,b)]\<bullet>x \<noteq> x"  | 
|
1241  | 
proof  | 
|
1242  | 
assume neg: "[(a,b)]\<bullet>x = x"  | 
|
1243  | 
from a1 have a1':"a\<in>(supp x)" by (simp add: fresh_def)  | 
|
1244  | 
from a2 have a2':"b\<notin>(supp x)" by (simp add: fresh_def)  | 
|
1245  | 
from a1' a2' have a3: "a\<noteq>b" by force  | 
|
1246  | 
from a1' have "([(a,b)]\<bullet>a)\<in>([(a,b)]\<bullet>(supp x))"  | 
|
1247  | 
by (simp only: pt_set_bij[OF at_pt_inst[OF at], OF at])  | 
|
1248  | 
hence "b\<in>([(a,b)]\<bullet>(supp x))" by (simp add: at_append[OF at] at_calc[OF at])  | 
|
1249  | 
hence "b\<in>(supp ([(a,b)]\<bullet>x))" by (simp add: pt_perm_supp[OF pt,OF at])  | 
|
1250  | 
with a2' neg show False by simp  | 
|
1251  | 
qed  | 
|
1252  | 
||
1253  | 
-- "three helper lemmas for the perm_fresh_fresh-lemma"  | 
|
1254  | 
lemma comprehension_neg_UNIV: "{b. \<not> P b} = UNIV - {b. P b}"
 | 
|
1255  | 
by (auto)  | 
|
1256  | 
||
1257  | 
lemma infinite_or_neg_infinite:  | 
|
1258  | 
assumes h:"infinite (UNIV::'a set)"  | 
|
1259  | 
  shows "infinite {b::'a. P b} \<or> infinite {b::'a. \<not> P b}"
 | 
|
1260  | 
proof (subst comprehension_neg_UNIV, case_tac "finite {b. P b}")
 | 
|
1261  | 
  assume j:"finite {b::'a. P b}"
 | 
|
1262  | 
  have "infinite ((UNIV::'a set) - {b::'a. P b})"
 | 
|
1263  | 
using Diff_infinite_finite[OF j h] by auto  | 
|
1264  | 
  thus "infinite {b::'a. P b} \<or> infinite (UNIV - {b::'a. P b})" ..
 | 
|
1265  | 
next  | 
|
1266  | 
  assume j:"infinite {b::'a. P b}"
 | 
|
1267  | 
  thus "infinite {b::'a. P b} \<or> infinite (UNIV - {b::'a. P b})" by simp
 | 
|
1268  | 
qed  | 
|
1269  | 
||
1270  | 
--"the co-set of a finite set is infinte"  | 
|
1271  | 
lemma finite_infinite:  | 
|
1272  | 
  assumes a: "finite {b::'x. P b}"
 | 
|
1273  | 
and b: "infinite (UNIV::'x set)"  | 
|
1274  | 
  shows "infinite {b. \<not>P b}"
 | 
|
1275  | 
using a and infinite_or_neg_infinite[OF b] by simp  | 
|
1276  | 
||
1277  | 
lemma pt_fresh_fresh:  | 
|
1278  | 
fixes x :: "'a"  | 
|
1279  | 
and a :: "'x"  | 
|
1280  | 
and b :: "'x"  | 
|
1281  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1282  | 
  and     at: "at TYPE ('x)"
 | 
|
1283  | 
and a1: "a\<sharp>x" and a2: "b\<sharp>x"  | 
|
1284  | 
shows "[(a,b)]\<bullet>x=x"  | 
|
1285  | 
proof (cases "a=b")  | 
|
1286  | 
assume c1: "a=b"  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
1287  | 
have "[(a,a)] \<triangleq> []" by (rule at_ds1[OF at])  | 
| 
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
1288  | 
hence "[(a,b)] \<triangleq> []" using c1 by simp  | 
| 17870 | 1289  | 
hence "[(a,b)]\<bullet>x=([]::'x prm)\<bullet>x" by (rule pt3[OF pt])  | 
1290  | 
thus ?thesis by (simp only: pt1[OF pt])  | 
|
1291  | 
next  | 
|
1292  | 
assume c2: "a\<noteq>b"  | 
|
1293  | 
  from a1 have f1: "finite {c. [(a,c)]\<bullet>x \<noteq> x}" by (simp add: fresh_def supp_def)
 | 
|
1294  | 
  from a2 have f2: "finite {c. [(b,c)]\<bullet>x \<noteq> x}" by (simp add: fresh_def supp_def)
 | 
|
1295  | 
  from f1 and f2 have f3: "finite {c. perm [(a,c)] x \<noteq> x \<or> perm [(b,c)] x \<noteq> x}" 
 | 
|
1296  | 
by (force simp only: Collect_disj_eq)  | 
|
1297  | 
  have "infinite {c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}" 
 | 
|
1298  | 
by (simp add: finite_infinite[OF f3,OF at4[OF at], simplified])  | 
|
1299  | 
  hence "infinite ({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b})" 
 | 
|
1300  | 
by (force dest: Diff_infinite_finite)  | 
|
1301  | 
  hence "({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b}) \<noteq> {}" 
 | 
|
1302  | 
by (auto iff del: finite_Diff_insert Diff_eq_empty_iff)  | 
|
1303  | 
  hence "\<exists>c. c\<in>({c. [(a,c)]\<bullet>x = x \<and> [(b,c)]\<bullet>x = x}-{a,b})" by (force)
 | 
|
1304  | 
then obtain c  | 
|
1305  | 
where eq1: "[(a,c)]\<bullet>x = x"  | 
|
1306  | 
and eq2: "[(b,c)]\<bullet>x = x"  | 
|
1307  | 
and ineq: "a\<noteq>c \<and> b\<noteq>c"  | 
|
1308  | 
by (force)  | 
|
1309  | 
hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>x)) = x" by simp  | 
|
1310  | 
hence eq3: "[(a,c),(b,c),(a,c)]\<bullet>x = x" by (simp add: pt2[OF pt,symmetric])  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
1311  | 
from c2 ineq have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" by (simp add: at_ds3[OF at])  | 
| 17870 | 1312  | 
hence "[(a,c),(b,c),(a,c)]\<bullet>x = [(a,b)]\<bullet>x" by (rule pt3[OF pt])  | 
1313  | 
thus ?thesis using eq3 by simp  | 
|
1314  | 
qed  | 
|
1315  | 
||
1316  | 
lemma pt_perm_compose:  | 
|
1317  | 
fixes pi1 :: "'x prm"  | 
|
1318  | 
and pi2 :: "'x prm"  | 
|
1319  | 
and x :: "'a"  | 
|
1320  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1321  | 
  and     at: "at TYPE('x)"
 | 
|
1322  | 
shows "pi2\<bullet>(pi1\<bullet>x) = (pi2\<bullet>pi1)\<bullet>(pi2\<bullet>x)"  | 
|
1323  | 
proof -  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
1324  | 
have "(pi2@pi1) \<triangleq> ((pi2\<bullet>pi1)@pi2)" by (rule at_ds8)  | 
| 17870 | 1325  | 
hence "(pi2@pi1)\<bullet>x = ((pi2\<bullet>pi1)@pi2)\<bullet>x" by (rule pt3[OF pt])  | 
1326  | 
thus ?thesis by (simp add: pt2[OF pt])  | 
|
1327  | 
qed  | 
|
1328  | 
||
1329  | 
lemma pt_perm_compose_rev:  | 
|
1330  | 
fixes pi1 :: "'x prm"  | 
|
1331  | 
and pi2 :: "'x prm"  | 
|
1332  | 
and x :: "'a"  | 
|
1333  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1334  | 
  and     at: "at TYPE('x)"
 | 
|
1335  | 
shows "(rev pi2)\<bullet>((rev pi1)\<bullet>x) = (rev pi1)\<bullet>(rev (pi1\<bullet>pi2)\<bullet>x)"  | 
|
1336  | 
proof -  | 
|
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
1337  | 
have "((rev pi2)@(rev pi1)) \<triangleq> ((rev pi1)@(rev (pi1\<bullet>pi2)))" by (rule at_ds9[OF at])  | 
| 17870 | 1338  | 
hence "((rev pi2)@(rev pi1))\<bullet>x = ((rev pi1)@(rev (pi1\<bullet>pi2)))\<bullet>x" by (rule pt3[OF pt])  | 
1339  | 
thus ?thesis by (simp add: pt2[OF pt])  | 
|
1340  | 
qed  | 
|
1341  | 
||
1342  | 
section {* facts about supports *}
 | 
|
1343  | 
(*==============================*)  | 
|
1344  | 
||
1345  | 
lemma supports_subset:  | 
|
1346  | 
fixes x :: "'a"  | 
|
1347  | 
and S1 :: "'x set"  | 
|
1348  | 
and S2 :: "'x set"  | 
|
1349  | 
assumes a: "S1 supports x"  | 
|
| 
18053
 
2719a6b7d95e
some minor tweaks in some proofs (nothing extraordinary)
 
urbanc 
parents: 
18048 
diff
changeset
 | 
1350  | 
and b: "S1 \<subseteq> S2"  | 
| 17870 | 1351  | 
shows "S2 supports x"  | 
1352  | 
using a b  | 
|
1353  | 
by (force simp add: "op supports_def")  | 
|
1354  | 
||
1355  | 
lemma supp_is_subset:  | 
|
1356  | 
fixes S :: "'x set"  | 
|
1357  | 
and x :: "'a"  | 
|
1358  | 
assumes a1: "S supports x"  | 
|
1359  | 
and a2: "finite S"  | 
|
1360  | 
shows "(supp x)\<subseteq>S"  | 
|
1361  | 
proof (rule ccontr)  | 
|
1362  | 
assume "\<not>(supp x \<subseteq> S)"  | 
|
1363  | 
hence "\<exists>a. a\<in>(supp x) \<and> a\<notin>S" by force  | 
|
1364  | 
then obtain a where b1: "a\<in>supp x" and b2: "a\<notin>S" by force  | 
|
1365  | 
from a1 b2 have "\<forall>b. (b\<notin>S \<longrightarrow> ([(a,b)]\<bullet>x = x))" by (unfold "op supports_def", force)  | 
|
1366  | 
  with a1 have "{b. [(a,b)]\<bullet>x \<noteq> x}\<subseteq>S" by (unfold "op supports_def", force)
 | 
|
1367  | 
  with a2 have "finite {b. [(a,b)]\<bullet>x \<noteq> x}" by (simp add: finite_subset)
 | 
|
1368  | 
hence "a\<notin>(supp x)" by (unfold supp_def, auto)  | 
|
1369  | 
with b1 show False by simp  | 
|
1370  | 
qed  | 
|
1371  | 
||
| 18264 | 1372  | 
lemma supp_supports:  | 
1373  | 
fixes x :: "'a"  | 
|
1374  | 
  assumes  pt: "pt TYPE('a) TYPE('x)"
 | 
|
1375  | 
  and      at: "at TYPE ('x)"
 | 
|
1376  | 
shows "((supp x)::'x set) supports x"  | 
|
1377  | 
proof (unfold "op supports_def", intro strip)  | 
|
1378  | 
fix a b  | 
|
1379  | 
assume "(a::'x)\<notin>(supp x) \<and> (b::'x)\<notin>(supp x)"  | 
|
1380  | 
hence "a\<sharp>x" and "b\<sharp>x" by (auto simp add: fresh_def)  | 
|
1381  | 
thus "[(a,b)]\<bullet>x = x" by (rule pt_fresh_fresh[OF pt, OF at])  | 
|
1382  | 
qed  | 
|
1383  | 
||
| 17870 | 1384  | 
lemma supports_finite:  | 
1385  | 
fixes S :: "'x set"  | 
|
1386  | 
and x :: "'a"  | 
|
1387  | 
assumes a1: "S supports x"  | 
|
1388  | 
and a2: "finite S"  | 
|
1389  | 
shows "finite ((supp x)::'x set)"  | 
|
1390  | 
proof -  | 
|
1391  | 
have "(supp x)\<subseteq>S" using a1 a2 by (rule supp_is_subset)  | 
|
1392  | 
thus ?thesis using a2 by (simp add: finite_subset)  | 
|
1393  | 
qed  | 
|
1394  | 
||
1395  | 
lemma supp_is_inter:  | 
|
1396  | 
fixes x :: "'a"  | 
|
1397  | 
  assumes  pt: "pt TYPE('a) TYPE('x)"
 | 
|
1398  | 
  and      at: "at TYPE ('x)"
 | 
|
1399  | 
  and      fs: "fs TYPE('a) TYPE('x)"
 | 
|
1400  | 
  shows "((supp x)::'x set) = (\<Inter> {S. finite S \<and> S supports x})"
 | 
|
1401  | 
proof (rule equalityI)  | 
|
1402  | 
  show "((supp x)::'x set) \<subseteq> (\<Inter> {S. finite S \<and> S supports x})"
 | 
|
1403  | 
proof (clarify)  | 
|
1404  | 
fix S c  | 
|
1405  | 
assume b: "c\<in>((supp x)::'x set)" and "finite (S::'x set)" and "S supports x"  | 
|
1406  | 
hence "((supp x)::'x set)\<subseteq>S" by (simp add: supp_is_subset)  | 
|
1407  | 
with b show "c\<in>S" by force  | 
|
1408  | 
qed  | 
|
1409  | 
next  | 
|
1410  | 
  show "(\<Inter> {S. finite S \<and> S supports x}) \<subseteq> ((supp x)::'x set)"
 | 
|
1411  | 
proof (clarify, simp)  | 
|
1412  | 
fix c  | 
|
1413  | 
assume d: "\<forall>(S::'x set). finite S \<and> S supports x \<longrightarrow> c\<in>S"  | 
|
1414  | 
have "((supp x)::'x set) supports x" by (rule supp_supports[OF pt, OF at])  | 
|
1415  | 
with d fs1[OF fs] show "c\<in>supp x" by force  | 
|
1416  | 
qed  | 
|
1417  | 
qed  | 
|
1418  | 
||
1419  | 
lemma supp_is_least_supports:  | 
|
1420  | 
fixes S :: "'x set"  | 
|
1421  | 
and x :: "'a"  | 
|
1422  | 
  assumes  pt: "pt TYPE('a) TYPE('x)"
 | 
|
1423  | 
  and      at: "at TYPE ('x)"
 | 
|
1424  | 
and a1: "S supports x"  | 
|
1425  | 
and a2: "finite S"  | 
|
1426  | 
and a3: "\<forall>S'. (finite S' \<and> S' supports x) \<longrightarrow> S\<subseteq>S'"  | 
|
1427  | 
shows "S = (supp x)"  | 
|
1428  | 
proof (rule equalityI)  | 
|
1429  | 
show "((supp x)::'x set)\<subseteq>S" using a1 a2 by (rule supp_is_subset)  | 
|
1430  | 
next  | 
|
1431  | 
have s1: "((supp x)::'x set) supports x" by (rule supp_supports[OF pt, OF at])  | 
|
1432  | 
have "((supp x)::'x set)\<subseteq>S" using a1 a2 by (rule supp_is_subset)  | 
|
1433  | 
hence "finite ((supp x)::'x set)" using a2 by (simp add: finite_subset)  | 
|
1434  | 
with s1 a3 show "S\<subseteq>supp x" by force  | 
|
1435  | 
qed  | 
|
1436  | 
||
1437  | 
lemma supports_set:  | 
|
1438  | 
fixes S :: "'x set"  | 
|
1439  | 
and X :: "'a set"  | 
|
1440  | 
  assumes  pt: "pt TYPE('a) TYPE('x)"
 | 
|
1441  | 
  and      at: "at TYPE ('x)"
 | 
|
1442  | 
and a: "\<forall>x\<in>X. (\<forall>(a::'x) (b::'x). a\<notin>S\<and>b\<notin>S \<longrightarrow> ([(a,b)]\<bullet>x)\<in>X)"  | 
|
1443  | 
shows "S supports X"  | 
|
1444  | 
using a  | 
|
1445  | 
apply(auto simp add: "op supports_def")  | 
|
1446  | 
apply(simp add: pt_set_bij1a[OF pt, OF at])  | 
|
1447  | 
apply(force simp add: pt_swap_bij[OF pt, OF at])  | 
|
1448  | 
apply(simp add: pt_set_bij1a[OF pt, OF at])  | 
|
1449  | 
done  | 
|
1450  | 
||
1451  | 
lemma supports_fresh:  | 
|
1452  | 
fixes S :: "'x set"  | 
|
1453  | 
and a :: "'x"  | 
|
1454  | 
and x :: "'a"  | 
|
1455  | 
assumes a1: "S supports x"  | 
|
1456  | 
and a2: "finite S"  | 
|
1457  | 
and a3: "a\<notin>S"  | 
|
1458  | 
shows "a\<sharp>x"  | 
|
1459  | 
proof (simp add: fresh_def)  | 
|
1460  | 
have "(supp x)\<subseteq>S" using a1 a2 by (rule supp_is_subset)  | 
|
1461  | 
thus "a\<notin>(supp x)" using a3 by force  | 
|
1462  | 
qed  | 
|
1463  | 
||
1464  | 
lemma at_fin_set_supports:  | 
|
1465  | 
fixes X::"'x set"  | 
|
1466  | 
  assumes at: "at TYPE('x)"
 | 
|
1467  | 
shows "X supports X"  | 
|
1468  | 
proof (simp add: "op supports_def", intro strip)  | 
|
1469  | 
fix a b  | 
|
1470  | 
assume "a\<notin>X \<and> b\<notin>X"  | 
|
1471  | 
thus "[(a,b)]\<bullet>X = X" by (force simp add: perm_set_def at_calc[OF at])  | 
|
1472  | 
qed  | 
|
1473  | 
||
1474  | 
lemma at_fin_set_supp:  | 
|
1475  | 
fixes X::"'x set"  | 
|
1476  | 
  assumes at: "at TYPE('x)"
 | 
|
1477  | 
and fs: "finite X"  | 
|
1478  | 
shows "(supp X) = X"  | 
|
1479  | 
proof -  | 
|
1480  | 
  have pt_set: "pt TYPE('x set) TYPE('x)" 
 | 
|
1481  | 
by (rule pt_set_inst[OF at_pt_inst[OF at]])  | 
|
1482  | 
have X_supports_X: "X supports X" by (rule at_fin_set_supports[OF at])  | 
|
1483  | 
show ?thesis using pt_set at X_supports_X fs  | 
|
1484  | 
proof (rule supp_is_least_supports[symmetric])  | 
|
1485  | 
show "\<forall>S'. finite S' \<and> S' supports X \<longrightarrow> X \<subseteq> S'"  | 
|
1486  | 
proof (auto)  | 
|
1487  | 
fix S'::"'x set" and x::"'x"  | 
|
1488  | 
assume f: "finite S'"  | 
|
1489  | 
and s: "S' supports X"  | 
|
1490  | 
and e1: "x\<in>X"  | 
|
1491  | 
show "x\<in>S'"  | 
|
1492  | 
proof (rule ccontr)  | 
|
1493  | 
assume e2: "x\<notin>S'"  | 
|
1494  | 
have "\<exists>b. b\<notin>(X\<union>S')" by (force intro: ex_in_inf[OF at] simp only: fs f)  | 
|
1495  | 
then obtain b where b1: "b\<notin>X" and b2: "b\<notin>S'" by (auto)  | 
|
1496  | 
from s e2 b2 have c1: "[(x,b)]\<bullet>X=X" by (simp add: "op supports_def")  | 
|
1497  | 
from e1 b1 have c2: "[(x,b)]\<bullet>X\<noteq>X" by (force simp add: perm_set_def at_calc[OF at])  | 
|
1498  | 
show "False" using c1 c2 by simp  | 
|
1499  | 
qed  | 
|
1500  | 
qed  | 
|
1501  | 
qed  | 
|
1502  | 
qed  | 
|
1503  | 
||
1504  | 
section {* Permutations acting on Functions *}
 | 
|
1505  | 
(*==========================================*)  | 
|
1506  | 
||
1507  | 
lemma pt_fun_app_eq:  | 
|
1508  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1509  | 
and x :: "'a"  | 
|
1510  | 
and pi :: "'x prm"  | 
|
1511  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1512  | 
  and     at: "at TYPE('x)"
 | 
|
1513  | 
shows "pi\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)"  | 
|
1514  | 
by (simp add: perm_fun_def pt_rev_pi[OF pt, OF at])  | 
|
1515  | 
||
1516  | 
||
1517  | 
--"sometimes pt_fun_app_eq does to much; this lemma 'corrects it'"  | 
|
1518  | 
lemma pt_perm:  | 
|
1519  | 
fixes x :: "'a"  | 
|
1520  | 
and pi1 :: "'x prm"  | 
|
1521  | 
and pi2 :: "'x prm"  | 
|
1522  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1523  | 
  and     at: "at TYPE ('x)"
 | 
|
1524  | 
shows "(pi1\<bullet>perm pi2)(pi1\<bullet>x) = pi1\<bullet>(pi2\<bullet>x)"  | 
|
1525  | 
by (simp add: pt_fun_app_eq[OF pt, OF at])  | 
|
1526  | 
||
1527  | 
||
1528  | 
lemma pt_fun_eq:  | 
|
1529  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1530  | 
and pi :: "'x prm"  | 
|
1531  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1532  | 
  and     at: "at TYPE('x)"
 | 
|
1533  | 
shows "(pi\<bullet>f = f) = (\<forall> x. pi\<bullet>(f x) = f (pi\<bullet>x))" (is "?LHS = ?RHS")  | 
|
1534  | 
proof  | 
|
1535  | 
assume a: "?LHS"  | 
|
1536  | 
show "?RHS"  | 
|
1537  | 
proof  | 
|
1538  | 
fix x  | 
|
1539  | 
have "pi\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)" by (simp add: pt_fun_app_eq[OF pt, OF at])  | 
|
1540  | 
also have "\<dots> = f (pi\<bullet>x)" using a by simp  | 
|
1541  | 
finally show "pi\<bullet>(f x) = f (pi\<bullet>x)" by simp  | 
|
1542  | 
qed  | 
|
1543  | 
next  | 
|
1544  | 
assume b: "?RHS"  | 
|
1545  | 
show "?LHS"  | 
|
1546  | 
proof (rule ccontr)  | 
|
1547  | 
assume "(pi\<bullet>f) \<noteq> f"  | 
|
1548  | 
hence "\<exists>c. (pi\<bullet>f) c \<noteq> f c" by (simp add: expand_fun_eq)  | 
|
1549  | 
then obtain c where b1: "(pi\<bullet>f) c \<noteq> f c" by force  | 
|
1550  | 
from b have "pi\<bullet>(f ((rev pi)\<bullet>c)) = f (pi\<bullet>((rev pi)\<bullet>c))" by force  | 
|
1551  | 
hence "(pi\<bullet>f)(pi\<bullet>((rev pi)\<bullet>c)) = f (pi\<bullet>((rev pi)\<bullet>c))"  | 
|
1552  | 
by (simp add: pt_fun_app_eq[OF pt, OF at])  | 
|
1553  | 
hence "(pi\<bullet>f) c = f c" by (simp add: pt_pi_rev[OF pt, OF at])  | 
|
1554  | 
with b1 show "False" by simp  | 
|
1555  | 
qed  | 
|
1556  | 
qed  | 
|
1557  | 
||
1558  | 
-- "two helper lemmas for the equivariance of functions"  | 
|
1559  | 
lemma pt_swap_eq_aux:  | 
|
1560  | 
fixes y :: "'a"  | 
|
1561  | 
and pi :: "'x prm"  | 
|
1562  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1563  | 
and a: "\<forall>(a::'x) (b::'x). [(a,b)]\<bullet>y = y"  | 
|
1564  | 
shows "pi\<bullet>y = y"  | 
|
1565  | 
proof(induct pi)  | 
|
1566  | 
case Nil show ?case by (simp add: pt1[OF pt])  | 
|
1567  | 
next  | 
|
1568  | 
case (Cons x xs)  | 
|
1569  | 
have "\<exists>a b. x=(a,b)" by force  | 
|
1570  | 
then obtain a b where p: "x=(a,b)" by force  | 
|
1571  | 
assume i: "xs\<bullet>y = y"  | 
|
1572  | 
have "x#xs = [x]@xs" by simp  | 
|
1573  | 
hence "(x#xs)\<bullet>y = ([x]@xs)\<bullet>y" by simp  | 
|
1574  | 
hence "(x#xs)\<bullet>y = [x]\<bullet>(xs\<bullet>y)" by (simp only: pt2[OF pt])  | 
|
| 18264 | 1575  | 
thus ?case using a i p by force  | 
| 17870 | 1576  | 
qed  | 
1577  | 
||
1578  | 
lemma pt_swap_eq:  | 
|
1579  | 
fixes y :: "'a"  | 
|
1580  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1581  | 
shows "(\<forall>(a::'x) (b::'x). [(a,b)]\<bullet>y = y) = (\<forall>pi::'x prm. pi\<bullet>y = y)"  | 
|
1582  | 
by (force intro: pt_swap_eq_aux[OF pt])  | 
|
1583  | 
||
1584  | 
lemma pt_eqvt_fun1a:  | 
|
1585  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1586  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1587  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1588  | 
  and     at:  "at TYPE('x)"
 | 
|
1589  | 
  and     a:   "((supp f)::'x set)={}"
 | 
|
1590  | 
shows "\<forall>(pi::'x prm). pi\<bullet>f = f"  | 
|
1591  | 
proof (intro strip)  | 
|
1592  | 
fix pi  | 
|
1593  | 
have "\<forall>a b. a\<notin>((supp f)::'x set) \<and> b\<notin>((supp f)::'x set) \<longrightarrow> (([(a,b)]\<bullet>f) = f)"  | 
|
1594  | 
by (intro strip, fold fresh_def,  | 
|
1595  | 
simp add: pt_fresh_fresh[OF pt_fun_inst[OF pta, OF ptb, OF at],OF at])  | 
|
1596  | 
with a have "\<forall>(a::'x) (b::'x). ([(a,b)]\<bullet>f) = f" by force  | 
|
1597  | 
hence "\<forall>(pi::'x prm). pi\<bullet>f = f"  | 
|
1598  | 
by (simp add: pt_swap_eq[OF pt_fun_inst[OF pta, OF ptb, OF at]])  | 
|
1599  | 
thus "(pi::'x prm)\<bullet>f = f" by simp  | 
|
1600  | 
qed  | 
|
1601  | 
||
1602  | 
lemma pt_eqvt_fun1b:  | 
|
1603  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1604  | 
assumes a: "\<forall>(pi::'x prm). pi\<bullet>f = f"  | 
|
1605  | 
  shows "((supp f)::'x set)={}"
 | 
|
1606  | 
using a by (simp add: supp_def)  | 
|
1607  | 
||
1608  | 
lemma pt_eqvt_fun1:  | 
|
1609  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1610  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1611  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1612  | 
  and     at: "at TYPE('x)"
 | 
|
1613  | 
  shows "(((supp f)::'x set)={}) = (\<forall>(pi::'x prm). pi\<bullet>f = f)" (is "?LHS = ?RHS")
 | 
|
1614  | 
by (rule iffI, simp add: pt_eqvt_fun1a[OF pta, OF ptb, OF at], simp add: pt_eqvt_fun1b)  | 
|
1615  | 
||
1616  | 
lemma pt_eqvt_fun2a:  | 
|
1617  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1618  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1619  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1620  | 
  and     at: "at TYPE('x)"
 | 
|
1621  | 
  assumes a: "((supp f)::'x set)={}"
 | 
|
1622  | 
shows "\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x)"  | 
|
1623  | 
proof (intro strip)  | 
|
1624  | 
fix pi x  | 
|
1625  | 
from a have b: "\<forall>(pi::'x prm). pi\<bullet>f = f" by (simp add: pt_eqvt_fun1[OF pta, OF ptb, OF at])  | 
|
1626  | 
have "(pi::'x prm)\<bullet>(f x) = (pi\<bullet>f)(pi\<bullet>x)" by (simp add: pt_fun_app_eq[OF pta, OF at])  | 
|
1627  | 
with b show "(pi::'x prm)\<bullet>(f x) = f (pi\<bullet>x)" by force  | 
|
1628  | 
qed  | 
|
1629  | 
||
1630  | 
lemma pt_eqvt_fun2b:  | 
|
1631  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1632  | 
  assumes pt1: "pt TYPE('a) TYPE('x)"
 | 
|
1633  | 
  and     pt2: "pt TYPE('b) TYPE('x)"
 | 
|
1634  | 
  and     at: "at TYPE('x)"
 | 
|
1635  | 
assumes a: "\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x)"  | 
|
1636  | 
  shows "((supp f)::'x set)={}"
 | 
|
1637  | 
proof -  | 
|
1638  | 
from a have "\<forall>(pi::'x prm). pi\<bullet>f = f" by (simp add: pt_fun_eq[OF pt1, OF at, symmetric])  | 
|
1639  | 
thus ?thesis by (simp add: supp_def)  | 
|
1640  | 
qed  | 
|
1641  | 
||
1642  | 
lemma pt_eqvt_fun2:  | 
|
1643  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1644  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1645  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1646  | 
  and     at: "at TYPE('x)"
 | 
|
1647  | 
  shows "(((supp f)::'x set)={}) = (\<forall>(pi::'x prm) (x::'a). pi\<bullet>(f x) = f(pi\<bullet>x))" 
 | 
|
1648  | 
by (rule iffI,  | 
|
1649  | 
simp add: pt_eqvt_fun2a[OF pta, OF ptb, OF at],  | 
|
1650  | 
simp add: pt_eqvt_fun2b[OF pta, OF ptb, OF at])  | 
|
1651  | 
||
1652  | 
lemma pt_supp_fun_subset:  | 
|
1653  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1654  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1655  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1656  | 
  and     at: "at TYPE('x)" 
 | 
|
1657  | 
and f1: "finite ((supp f)::'x set)"  | 
|
1658  | 
and f2: "finite ((supp x)::'x set)"  | 
|
1659  | 
shows "supp (f x) \<subseteq> (((supp f)\<union>(supp x))::'x set)"  | 
|
1660  | 
proof -  | 
|
1661  | 
have s1: "((supp f)\<union>((supp x)::'x set)) supports (f x)"  | 
|
1662  | 
proof (simp add: "op supports_def", fold fresh_def, auto)  | 
|
1663  | 
fix a::"'x" and b::"'x"  | 
|
1664  | 
assume "a\<sharp>f" and "b\<sharp>f"  | 
|
1665  | 
hence a1: "[(a,b)]\<bullet>f = f"  | 
|
1666  | 
by (rule pt_fresh_fresh[OF pt_fun_inst[OF pta, OF ptb, OF at], OF at])  | 
|
1667  | 
assume "a\<sharp>x" and "b\<sharp>x"  | 
|
1668  | 
hence a2: "[(a,b)]\<bullet>x = x" by (rule pt_fresh_fresh[OF pta, OF at])  | 
|
1669  | 
from a1 a2 show "[(a,b)]\<bullet>(f x) = (f x)" by (simp add: pt_fun_app_eq[OF pta, OF at])  | 
|
1670  | 
qed  | 
|
1671  | 
from f1 f2 have "finite ((supp f)\<union>((supp x)::'x set))" by force  | 
|
1672  | 
with s1 show ?thesis by (rule supp_is_subset)  | 
|
1673  | 
qed  | 
|
1674  | 
||
1675  | 
lemma pt_empty_supp_fun_subset:  | 
|
1676  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
1677  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1678  | 
  and     ptb: "pt TYPE('b) TYPE('x)"
 | 
|
1679  | 
  and     at:  "at TYPE('x)" 
 | 
|
1680  | 
  and     e:   "(supp f)=({}::'x set)"
 | 
|
1681  | 
shows "supp (f x) \<subseteq> ((supp x)::'x set)"  | 
|
1682  | 
proof (unfold supp_def, auto)  | 
|
1683  | 
fix a::"'x"  | 
|
1684  | 
  assume a1: "finite {b. [(a, b)]\<bullet>x \<noteq> x}"
 | 
|
1685  | 
  assume "infinite {b. [(a, b)]\<bullet>(f x) \<noteq> f x}"
 | 
|
1686  | 
  hence a2: "infinite {b. f ([(a, b)]\<bullet>x) \<noteq> f x}" using e
 | 
|
1687  | 
by (simp add: pt_eqvt_fun2[OF pta, OF ptb, OF at])  | 
|
1688  | 
  have a3: "{b. f ([(a,b)]\<bullet>x) \<noteq> f x}\<subseteq>{b. [(a,b)]\<bullet>x \<noteq> x}" by force
 | 
|
1689  | 
from a1 a2 a3 show False by (force dest: finite_subset)  | 
|
1690  | 
qed  | 
|
1691  | 
||
| 18264 | 1692  | 
section {* Facts about the support of finite sets of finitely supported things *}
 | 
1693  | 
(*=============================================================================*)  | 
|
1694  | 
||
1695  | 
constdefs  | 
|
1696  | 
  X_to_Un_supp :: "('a set) \<Rightarrow> 'x set"
 | 
|
1697  | 
"X_to_Un_supp X \<equiv> \<Union>x\<in>X. ((supp x)::'x set)"  | 
|
1698  | 
||
1699  | 
lemma UNION_f_eqvt:  | 
|
1700  | 
  fixes X::"('a set)"
 | 
|
1701  | 
and f::"'a \<Rightarrow> 'x set"  | 
|
1702  | 
and pi::"'x prm"  | 
|
1703  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1704  | 
  and     at: "at TYPE('x)"
 | 
|
1705  | 
shows "pi\<bullet>(\<Union>x\<in>X. f x) = (\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x)"  | 
|
1706  | 
proof -  | 
|
1707  | 
  have pt_x: "pt TYPE('x) TYPE('x)" by (force intro: at_pt_inst at)
 | 
|
1708  | 
show ?thesis  | 
|
| 18351 | 1709  | 
proof (rule equalityI)  | 
1710  | 
case goal1  | 
|
1711  | 
show "pi\<bullet>(\<Union>x\<in>X. f x) \<subseteq> (\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x)"  | 
|
1712  | 
apply(auto simp add: perm_set_def)  | 
|
1713  | 
apply(rule_tac x="pi\<bullet>xa" in exI)  | 
|
1714  | 
apply(rule conjI)  | 
|
1715  | 
apply(rule_tac x="xa" in exI)  | 
|
1716  | 
apply(simp)  | 
|
1717  | 
apply(subgoal_tac "(pi\<bullet>f) (pi\<bullet>xa) = pi\<bullet>(f xa)")(*A*)  | 
|
1718  | 
apply(simp)  | 
|
1719  | 
apply(rule pt_set_bij2[OF pt_x, OF at])  | 
|
1720  | 
apply(assumption)  | 
|
1721  | 
(*A*)  | 
|
1722  | 
apply(rule sym)  | 
|
1723  | 
apply(rule pt_fun_app_eq[OF pt, OF at])  | 
|
1724  | 
done  | 
|
1725  | 
next  | 
|
1726  | 
case goal2  | 
|
1727  | 
show "(\<Union>x\<in>(pi\<bullet>X). (pi\<bullet>f) x) \<subseteq> pi\<bullet>(\<Union>x\<in>X. f x)"  | 
|
1728  | 
apply(auto simp add: perm_set_def)  | 
|
1729  | 
apply(rule_tac x="(rev pi)\<bullet>x" in exI)  | 
|
1730  | 
apply(rule conjI)  | 
|
1731  | 
apply(simp add: pt_pi_rev[OF pt_x, OF at])  | 
|
1732  | 
apply(rule_tac x="a" in bexI)  | 
|
1733  | 
apply(simp add: pt_set_bij1[OF pt_x, OF at])  | 
|
1734  | 
apply(simp add: pt_fun_app_eq[OF pt, OF at])  | 
|
1735  | 
apply(assumption)  | 
|
1736  | 
done  | 
|
1737  | 
qed  | 
|
| 18264 | 1738  | 
qed  | 
1739  | 
||
1740  | 
lemma X_to_Un_supp_eqvt:  | 
|
1741  | 
  fixes X::"('a set)"
 | 
|
1742  | 
and pi::"'x prm"  | 
|
1743  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1744  | 
  and     at: "at TYPE('x)"
 | 
|
1745  | 
shows "pi\<bullet>(X_to_Un_supp X) = ((X_to_Un_supp (pi\<bullet>X))::'x set)"  | 
|
1746  | 
apply(simp add: X_to_Un_supp_def)  | 
|
1747  | 
apply(simp add: UNION_f_eqvt[OF pt, OF at] perm_fun_def)  | 
|
1748  | 
apply(simp add: pt_perm_supp[OF pt, OF at])  | 
|
1749  | 
apply(simp add: pt_pi_rev[OF pt, OF at])  | 
|
1750  | 
done  | 
|
1751  | 
||
1752  | 
lemma Union_supports_set:  | 
|
1753  | 
  fixes X::"('a set)"
 | 
|
1754  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1755  | 
  and     at: "at TYPE('x)"
 | 
|
1756  | 
shows "(\<Union>x\<in>X. ((supp x)::'x set)) supports X"  | 
|
1757  | 
apply(simp add: "op supports_def" fresh_def[symmetric])  | 
|
1758  | 
apply(rule allI)+  | 
|
1759  | 
apply(rule impI)  | 
|
1760  | 
apply(erule conjE)  | 
|
1761  | 
apply(simp add: perm_set_def)  | 
|
1762  | 
apply(auto)  | 
|
1763  | 
apply(subgoal_tac "[(a,b)]\<bullet>aa = aa")(*A*)  | 
|
1764  | 
apply(simp)  | 
|
1765  | 
apply(rule pt_fresh_fresh[OF pt, OF at])  | 
|
1766  | 
apply(force)  | 
|
1767  | 
apply(force)  | 
|
1768  | 
apply(rule_tac x="x" in exI)  | 
|
1769  | 
apply(simp)  | 
|
1770  | 
apply(rule sym)  | 
|
1771  | 
apply(rule pt_fresh_fresh[OF pt, OF at])  | 
|
1772  | 
apply(force)+  | 
|
1773  | 
done  | 
|
1774  | 
||
1775  | 
lemma Union_of_fin_supp_sets:  | 
|
1776  | 
  fixes X::"('a set)"
 | 
|
1777  | 
  assumes fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1778  | 
and fi: "finite X"  | 
|
1779  | 
shows "finite (\<Union>x\<in>X. ((supp x)::'x set))"  | 
|
1780  | 
using fi by (induct, auto simp add: fs1[OF fs])  | 
|
1781  | 
||
1782  | 
lemma Union_included_in_supp:  | 
|
1783  | 
  fixes X::"('a set)"
 | 
|
1784  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1785  | 
  and     at: "at TYPE('x)"
 | 
|
1786  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1787  | 
and fi: "finite X"  | 
|
1788  | 
shows "(\<Union>x\<in>X. ((supp x)::'x set)) \<subseteq> supp X"  | 
|
1789  | 
proof -  | 
|
1790  | 
have "supp ((X_to_Un_supp X)::'x set) \<subseteq> ((supp X)::'x set)"  | 
|
1791  | 
apply(rule pt_empty_supp_fun_subset)  | 
|
1792  | 
apply(force intro: pt_set_inst at_pt_inst pt at)+  | 
|
1793  | 
apply(rule pt_eqvt_fun2b)  | 
|
1794  | 
apply(force intro: pt_set_inst at_pt_inst pt at)+  | 
|
| 18351 | 1795  | 
apply(rule allI)+  | 
| 18264 | 1796  | 
apply(rule X_to_Un_supp_eqvt[OF pt, OF at])  | 
1797  | 
done  | 
|
1798  | 
hence "supp (\<Union>x\<in>X. ((supp x)::'x set)) \<subseteq> ((supp X)::'x set)" by (simp add: X_to_Un_supp_def)  | 
|
1799  | 
moreover  | 
|
1800  | 
have "supp (\<Union>x\<in>X. ((supp x)::'x set)) = (\<Union>x\<in>X. ((supp x)::'x set))"  | 
|
1801  | 
apply(rule at_fin_set_supp[OF at])  | 
|
1802  | 
apply(rule Union_of_fin_supp_sets[OF fs, OF fi])  | 
|
1803  | 
done  | 
|
1804  | 
ultimately show ?thesis by force  | 
|
1805  | 
qed  | 
|
1806  | 
||
1807  | 
lemma supp_of_fin_sets:  | 
|
1808  | 
  fixes X::"('a set)"
 | 
|
1809  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1810  | 
  and     at: "at TYPE('x)"
 | 
|
1811  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1812  | 
and fi: "finite X"  | 
|
1813  | 
shows "(supp X) = (\<Union>x\<in>X. ((supp x)::'x set))"  | 
|
| 18351 | 1814  | 
apply(rule equalityI)  | 
| 18264 | 1815  | 
apply(rule supp_is_subset)  | 
1816  | 
apply(rule Union_supports_set[OF pt, OF at])  | 
|
1817  | 
apply(rule Union_of_fin_supp_sets[OF fs, OF fi])  | 
|
1818  | 
apply(rule Union_included_in_supp[OF pt, OF at, OF fs, OF fi])  | 
|
1819  | 
done  | 
|
1820  | 
||
1821  | 
lemma supp_fin_union:  | 
|
1822  | 
  fixes X::"('a set)"
 | 
|
1823  | 
  and   Y::"('a set)"
 | 
|
1824  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1825  | 
  and     at: "at TYPE('x)"
 | 
|
1826  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1827  | 
and f1: "finite X"  | 
|
1828  | 
and f2: "finite Y"  | 
|
1829  | 
shows "(supp (X\<union>Y)) = (supp X)\<union>((supp Y)::'x set)"  | 
|
1830  | 
using f1 f2 by (force simp add: supp_of_fin_sets[OF pt, OF at, OF fs])  | 
|
1831  | 
||
1832  | 
lemma supp_fin_insert:  | 
|
1833  | 
  fixes X::"('a set)"
 | 
|
1834  | 
and x::"'a"  | 
|
1835  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1836  | 
  and     at: "at TYPE('x)"
 | 
|
1837  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1838  | 
and f: "finite X"  | 
|
1839  | 
shows "(supp (insert x X)) = (supp x)\<union>((supp X)::'x set)"  | 
|
1840  | 
proof -  | 
|
1841  | 
  have "(supp (insert x X)) = ((supp ({x}\<union>(X::'a set)))::'x set)" by simp
 | 
|
1842  | 
  also have "\<dots> = (supp {x})\<union>(supp X)"
 | 
|
1843  | 
by (rule supp_fin_union[OF pt, OF at, OF fs], simp_all add: f)  | 
|
1844  | 
finally show "(supp (insert x X)) = (supp x)\<union>((supp X)::'x set)"  | 
|
1845  | 
by (simp add: supp_singleton)  | 
|
1846  | 
qed  | 
|
1847  | 
||
1848  | 
lemma fresh_fin_union:  | 
|
1849  | 
  fixes X::"('a set)"
 | 
|
1850  | 
  and   Y::"('a set)"
 | 
|
1851  | 
and a::"'x"  | 
|
1852  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1853  | 
  and     at: "at TYPE('x)"
 | 
|
1854  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1855  | 
and f1: "finite X"  | 
|
1856  | 
and f2: "finite Y"  | 
|
1857  | 
shows "a\<sharp>(X\<union>Y) = (a\<sharp>X \<and> a\<sharp>Y)"  | 
|
1858  | 
apply(simp add: fresh_def)  | 
|
1859  | 
apply(simp add: supp_fin_union[OF pt, OF at, OF fs, OF f1, OF f2])  | 
|
1860  | 
done  | 
|
1861  | 
||
1862  | 
lemma fresh_fin_insert:  | 
|
1863  | 
  fixes X::"('a set)"
 | 
|
1864  | 
and x::"'a"  | 
|
1865  | 
and a::"'x"  | 
|
1866  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1867  | 
  and     at: "at TYPE('x)"
 | 
|
1868  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1869  | 
and f: "finite X"  | 
|
1870  | 
shows "a\<sharp>(insert x X) = (a\<sharp>x \<and> a\<sharp>X)"  | 
|
1871  | 
apply(simp add: fresh_def)  | 
|
1872  | 
apply(simp add: supp_fin_insert[OF pt, OF at, OF fs, OF f])  | 
|
1873  | 
done  | 
|
1874  | 
||
1875  | 
lemma fresh_fin_insert1:  | 
|
1876  | 
  fixes X::"('a set)"
 | 
|
1877  | 
and x::"'a"  | 
|
1878  | 
and a::"'x"  | 
|
1879  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1880  | 
  and     at: "at TYPE('x)"
 | 
|
1881  | 
  and     fs: "fs TYPE('a) TYPE('x)" 
 | 
|
1882  | 
and f: "finite X"  | 
|
1883  | 
and a1: "a\<sharp>x"  | 
|
1884  | 
and a2: "a\<sharp>X"  | 
|
1885  | 
shows "a\<sharp>(insert x X)"  | 
|
1886  | 
using a1 a2  | 
|
1887  | 
apply(simp add: fresh_fin_insert[OF pt, OF at, OF fs, OF f])  | 
|
1888  | 
done  | 
|
1889  | 
||
1890  | 
lemma pt_list_set_pi:  | 
|
1891  | 
fixes pi :: "'x prm"  | 
|
1892  | 
and xs :: "'a list"  | 
|
1893  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1894  | 
shows "pi\<bullet>(set xs) = set (pi\<bullet>xs)"  | 
|
1895  | 
by (induct xs, auto simp add: perm_set_def pt1[OF pt])  | 
|
1896  | 
||
1897  | 
lemma pt_list_set_supp:  | 
|
1898  | 
fixes xs :: "'a list"  | 
|
1899  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1900  | 
  and     at: "at TYPE('x)"
 | 
|
1901  | 
  and     fs: "fs TYPE('a) TYPE('x)"
 | 
|
1902  | 
shows "supp (set xs) = ((supp xs)::'x set)"  | 
|
1903  | 
proof -  | 
|
1904  | 
have "supp (set xs) = (\<Union>x\<in>(set xs). ((supp x)::'x set))"  | 
|
1905  | 
by (rule supp_of_fin_sets[OF pt, OF at, OF fs], rule finite_set)  | 
|
1906  | 
also have "(\<Union>x\<in>(set xs). ((supp x)::'x set)) = (supp xs)"  | 
|
1907  | 
proof(induct xs)  | 
|
1908  | 
case Nil show ?case by (simp add: supp_list_nil)  | 
|
1909  | 
next  | 
|
1910  | 
case (Cons h t) thus ?case by (simp add: supp_list_cons)  | 
|
1911  | 
qed  | 
|
1912  | 
finally show ?thesis by simp  | 
|
1913  | 
qed  | 
|
1914  | 
||
1915  | 
lemma pt_list_set_fresh:  | 
|
1916  | 
fixes a :: "'x"  | 
|
1917  | 
and xs :: "'a list"  | 
|
1918  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1919  | 
  and     at: "at TYPE('x)"
 | 
|
1920  | 
  and     fs: "fs TYPE('a) TYPE('x)"
 | 
|
1921  | 
and a: "a\<sharp>xs"  | 
|
1922  | 
shows "a\<sharp>(set xs) = a\<sharp>xs"  | 
|
1923  | 
by (simp add: fresh_def pt_list_set_supp[OF pt, OF at, OF fs])  | 
|
1924  | 
||
| 17870 | 1925  | 
section {* Andy's freshness lemma *}
 | 
1926  | 
(*================================*)  | 
|
1927  | 
||
1928  | 
lemma freshness_lemma:  | 
|
1929  | 
fixes h :: "'x\<Rightarrow>'a"  | 
|
1930  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
1931  | 
  and     at:  "at TYPE('x)" 
 | 
|
1932  | 
and f1: "finite ((supp h)::'x set)"  | 
|
1933  | 
and a: "\<exists>a::'x. (a\<sharp>h \<and> a\<sharp>(h a))"  | 
|
1934  | 
shows "\<exists>fr::'a. \<forall>a::'x. a\<sharp>h \<longrightarrow> (h a) = fr"  | 
|
1935  | 
proof -  | 
|
1936  | 
  have ptb: "pt TYPE('x) TYPE('x)" by (simp add: at_pt_inst[OF at]) 
 | 
|
1937  | 
  have ptc: "pt TYPE('x\<Rightarrow>'a) TYPE('x)" by (simp add: pt_fun_inst[OF ptb, OF pta, OF at]) 
 | 
|
1938  | 
from a obtain a0 where a1: "a0\<sharp>h" and a2: "a0\<sharp>(h a0)" by force  | 
|
1939  | 
show ?thesis  | 
|
1940  | 
proof  | 
|
1941  | 
let ?fr = "h (a0::'x)"  | 
|
1942  | 
show "\<forall>(a::'x). (a\<sharp>h \<longrightarrow> ((h a) = ?fr))"  | 
|
1943  | 
proof (intro strip)  | 
|
1944  | 
fix a  | 
|
1945  | 
assume a3: "(a::'x)\<sharp>h"  | 
|
1946  | 
show "h (a::'x) = h a0"  | 
|
1947  | 
proof (cases "a=a0")  | 
|
1948  | 
case True thus "h (a::'x) = h a0" by simp  | 
|
1949  | 
next  | 
|
1950  | 
case False  | 
|
1951  | 
assume "a\<noteq>a0"  | 
|
1952  | 
hence c1: "a\<notin>((supp a0)::'x set)" by (simp add: fresh_def[symmetric] at_fresh[OF at])  | 
|
1953  | 
have c2: "a\<notin>((supp h)::'x set)" using a3 by (simp add: fresh_def)  | 
|
1954  | 
from c1 c2 have c3: "a\<notin>((supp h)\<union>((supp a0)::'x set))" by force  | 
|
1955  | 
have f2: "finite ((supp a0)::'x set)" by (simp add: at_supp[OF at])  | 
|
1956  | 
from f1 f2 have "((supp (h a0))::'x set)\<subseteq>((supp h)\<union>(supp a0))"  | 
|
1957  | 
by (simp add: pt_supp_fun_subset[OF ptb, OF pta, OF at])  | 
|
1958  | 
hence "a\<notin>((supp (h a0))::'x set)" using c3 by force  | 
|
1959  | 
hence "a\<sharp>(h a0)" by (simp add: fresh_def)  | 
|
1960  | 
with a2 have d1: "[(a0,a)]\<bullet>(h a0) = (h a0)" by (rule pt_fresh_fresh[OF pta, OF at])  | 
|
1961  | 
from a1 a3 have d2: "[(a0,a)]\<bullet>h = h" by (rule pt_fresh_fresh[OF ptc, OF at])  | 
|
1962  | 
from d1 have "h a0 = [(a0,a)]\<bullet>(h a0)" by simp  | 
|
1963  | 
also have "\<dots>= ([(a0,a)]\<bullet>h)([(a0,a)]\<bullet>a0)" by (simp add: pt_fun_app_eq[OF ptb, OF at])  | 
|
1964  | 
also have "\<dots> = h ([(a0,a)]\<bullet>a0)" using d2 by simp  | 
|
1965  | 
also have "\<dots> = h a" by (simp add: at_calc[OF at])  | 
|
1966  | 
finally show "h a = h a0" by simp  | 
|
1967  | 
qed  | 
|
1968  | 
qed  | 
|
1969  | 
qed  | 
|
1970  | 
qed  | 
|
1971  | 
||
1972  | 
lemma freshness_lemma_unique:  | 
|
1973  | 
fixes h :: "'x\<Rightarrow>'a"  | 
|
1974  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1975  | 
  and     at: "at TYPE('x)" 
 | 
|
1976  | 
and f1: "finite ((supp h)::'x set)"  | 
|
1977  | 
and a: "\<exists>(a::'x). (a\<sharp>h \<and> a\<sharp>(h a))"  | 
|
1978  | 
shows "\<exists>!(fr::'a). \<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr"  | 
|
1979  | 
proof  | 
|
1980  | 
from pt at f1 a show "\<exists>fr::'a. \<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr" by (simp add: freshness_lemma)  | 
|
1981  | 
next  | 
|
1982  | 
fix fr1 fr2  | 
|
1983  | 
assume b1: "\<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr1"  | 
|
1984  | 
assume b2: "\<forall>a::'x. a\<sharp>h \<longrightarrow> h a = fr2"  | 
|
1985  | 
from a obtain a where "(a::'x)\<sharp>h" by force  | 
|
1986  | 
with b1 b2 have "h a = fr1 \<and> h a = fr2" by force  | 
|
1987  | 
thus "fr1 = fr2" by force  | 
|
1988  | 
qed  | 
|
1989  | 
||
1990  | 
-- "packaging the freshness lemma into a function"  | 
|
1991  | 
constdefs  | 
|
1992  | 
  fresh_fun :: "('x\<Rightarrow>'a)\<Rightarrow>'a"
 | 
|
1993  | 
"fresh_fun (h) \<equiv> THE fr. (\<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr)"  | 
|
1994  | 
||
1995  | 
lemma fresh_fun_app:  | 
|
1996  | 
fixes h :: "'x\<Rightarrow>'a"  | 
|
1997  | 
and a :: "'x"  | 
|
1998  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
1999  | 
  and     at: "at TYPE('x)" 
 | 
|
2000  | 
and f1: "finite ((supp h)::'x set)"  | 
|
2001  | 
and a: "\<exists>(a::'x). (a\<sharp>h \<and> a\<sharp>(h a))"  | 
|
2002  | 
and b: "a\<sharp>h"  | 
|
2003  | 
shows "(fresh_fun h) = (h a)"  | 
|
2004  | 
proof (unfold fresh_fun_def, rule the_equality)  | 
|
2005  | 
show "\<forall>(a'::'x). a'\<sharp>h \<longrightarrow> h a' = h a"  | 
|
2006  | 
proof (intro strip)  | 
|
2007  | 
fix a'::"'x"  | 
|
2008  | 
assume c: "a'\<sharp>h"  | 
|
2009  | 
from pt at f1 a have "\<exists>(fr::'a). \<forall>(a::'x). a\<sharp>h \<longrightarrow> (h a) = fr" by (rule freshness_lemma)  | 
|
2010  | 
with b c show "h a' = h a" by force  | 
|
2011  | 
qed  | 
|
2012  | 
next  | 
|
2013  | 
fix fr::"'a"  | 
|
2014  | 
assume "\<forall>a. a\<sharp>h \<longrightarrow> h a = fr"  | 
|
2015  | 
with b show "fr = h a" by force  | 
|
2016  | 
qed  | 
|
2017  | 
||
2018  | 
||
2019  | 
lemma fresh_fun_supports:  | 
|
2020  | 
fixes h :: "'x\<Rightarrow>'a"  | 
|
2021  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2022  | 
  and     at: "at TYPE('x)" 
 | 
|
2023  | 
and f1: "finite ((supp h)::'x set)"  | 
|
2024  | 
and a: "\<exists>(a::'x). (a\<sharp>h \<and> a\<sharp>(h a))"  | 
|
2025  | 
shows "((supp h)::'x set) supports (fresh_fun h)"  | 
|
2026  | 
apply(simp add: "op supports_def")  | 
|
2027  | 
apply(fold fresh_def)  | 
|
2028  | 
apply(auto)  | 
|
2029  | 
apply(subgoal_tac "\<exists>(a''::'x). a''\<sharp>(h,a,b)")(*A*)  | 
|
2030  | 
apply(erule exE)  | 
|
2031  | 
apply(simp add: fresh_prod)  | 
|
2032  | 
apply(auto)  | 
|
2033  | 
apply(rotate_tac 2)  | 
|
2034  | 
apply(drule fresh_fun_app[OF pt, OF at, OF f1, OF a])  | 
|
2035  | 
apply(simp add: at_fresh[OF at])  | 
|
2036  | 
apply(simp add: pt_fun_app_eq[OF at_pt_inst[OF at], OF at])  | 
|
2037  | 
apply(auto simp add: at_calc[OF at])  | 
|
2038  | 
apply(subgoal_tac "[(a, b)]\<bullet>h = h")(*B*)  | 
|
2039  | 
apply(simp)  | 
|
2040  | 
(*B*)  | 
|
2041  | 
apply(rule pt_fresh_fresh[OF pt_fun_inst[OF at_pt_inst[OF at], OF pt], OF at, OF at])  | 
|
2042  | 
apply(assumption)+  | 
|
2043  | 
(*A*)  | 
|
2044  | 
apply(rule at_exists_fresh[OF at])  | 
|
2045  | 
apply(simp add: supp_prod)  | 
|
2046  | 
apply(simp add: f1 at_supp[OF at])  | 
|
2047  | 
done  | 
|
2048  | 
||
2049  | 
lemma fresh_fun_equiv:  | 
|
2050  | 
fixes h :: "'x\<Rightarrow>'a"  | 
|
2051  | 
and pi:: "'x prm"  | 
|
2052  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
2053  | 
  and     at:  "at TYPE('x)" 
 | 
|
2054  | 
and f1: "finite ((supp h)::'x set)"  | 
|
2055  | 
and a1: "\<exists>(a::'x). (a\<sharp>h \<and> a\<sharp>(h a))"  | 
|
2056  | 
shows "pi\<bullet>(fresh_fun h) = fresh_fun(pi\<bullet>h)" (is "?LHS = ?RHS")  | 
|
2057  | 
proof -  | 
|
2058  | 
  have ptb: "pt TYPE('x) TYPE('x)" by (simp add: at_pt_inst[OF at]) 
 | 
|
2059  | 
  have ptc: "pt TYPE('x\<Rightarrow>'a) TYPE('x)" by (simp add: pt_fun_inst[OF ptb, OF pta, OF at]) 
 | 
|
2060  | 
have f2: "finite ((supp (pi\<bullet>h))::'x set)"  | 
|
2061  | 
proof -  | 
|
2062  | 
from f1 have "finite (pi\<bullet>((supp h)::'x set))" by (simp add: pt_set_finite_ineq[OF ptb, OF at])  | 
|
2063  | 
thus ?thesis by (simp add: pt_perm_supp[OF ptc, OF at])  | 
|
2064  | 
qed  | 
|
2065  | 
from a1 obtain a' where c0: "a'\<sharp>h \<and> a'\<sharp>(h a')" by force  | 
|
2066  | 
hence c1: "a'\<sharp>h" and c2: "a'\<sharp>(h a')" by simp_all  | 
|
2067  | 
have c3: "(pi\<bullet>a')\<sharp>(pi\<bullet>h)" using c1 by (simp add: pt_fresh_bij[OF ptc, OF at])  | 
|
2068  | 
have c4: "(pi\<bullet>a')\<sharp>(pi\<bullet>h) (pi\<bullet>a')"  | 
|
2069  | 
proof -  | 
|
2070  | 
from c2 have "(pi\<bullet>a')\<sharp>(pi\<bullet>(h a'))" by (simp add: pt_fresh_bij[OF pta, OF at])  | 
|
2071  | 
thus ?thesis by (simp add: pt_fun_app_eq[OF ptb, OF at])  | 
|
2072  | 
qed  | 
|
2073  | 
have a2: "\<exists>(a::'x). (a\<sharp>(pi\<bullet>h) \<and> a\<sharp>((pi\<bullet>h) a))" using c3 c4 by force  | 
|
2074  | 
have d1: "?LHS = pi\<bullet>(h a')" using c1 a1 by (simp add: fresh_fun_app[OF pta, OF at, OF f1])  | 
|
2075  | 
have d2: "?RHS = (pi\<bullet>h) (pi\<bullet>a')" using c3 a2 by (simp add: fresh_fun_app[OF pta, OF at, OF f2])  | 
|
2076  | 
show ?thesis using d1 d2 by (simp add: pt_fun_app_eq[OF ptb, OF at])  | 
|
2077  | 
qed  | 
|
2078  | 
||
2079  | 
section {* disjointness properties *}
 | 
|
2080  | 
(*=================================*)  | 
|
2081  | 
lemma dj_perm_forget:  | 
|
2082  | 
fixes pi::"'y prm"  | 
|
2083  | 
and x ::"'x"  | 
|
2084  | 
  assumes dj: "disjoint TYPE('x) TYPE('y)"
 | 
|
2085  | 
shows "pi\<bullet>x=x"  | 
|
2086  | 
using dj by (simp add: disjoint_def)  | 
|
2087  | 
||
2088  | 
lemma dj_perm_perm_forget:  | 
|
2089  | 
fixes pi1::"'x prm"  | 
|
2090  | 
and pi2::"'y prm"  | 
|
2091  | 
  assumes dj: "disjoint TYPE('x) TYPE('y)"
 | 
|
2092  | 
shows "pi2\<bullet>pi1=pi1"  | 
|
2093  | 
using dj by (induct pi1, auto simp add: disjoint_def)  | 
|
2094  | 
||
2095  | 
lemma dj_cp:  | 
|
2096  | 
fixes pi1::"'x prm"  | 
|
2097  | 
and pi2::"'y prm"  | 
|
2098  | 
and x ::"'a"  | 
|
2099  | 
  assumes cp: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2100  | 
  and     dj: "disjoint TYPE('y) TYPE('x)"
 | 
|
2101  | 
shows "pi1\<bullet>(pi2\<bullet>x) = (pi2)\<bullet>(pi1\<bullet>x)"  | 
|
2102  | 
by (simp add: cp1[OF cp] dj_perm_perm_forget[OF dj])  | 
|
2103  | 
||
2104  | 
lemma dj_supp:  | 
|
2105  | 
fixes a::"'x"  | 
|
2106  | 
  assumes dj: "disjoint TYPE('x) TYPE('y)"
 | 
|
2107  | 
  shows "(supp a) = ({}::'y set)"
 | 
|
2108  | 
apply(simp add: supp_def dj_perm_forget[OF dj])  | 
|
2109  | 
done  | 
|
2110  | 
||
2111  | 
||
2112  | 
section {* composition instances *}
 | 
|
2113  | 
(* ============================= *)  | 
|
2114  | 
||
2115  | 
lemma cp_list_inst:  | 
|
2116  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2117  | 
  shows "cp TYPE ('a list) TYPE('x) TYPE('y)"
 | 
|
2118  | 
using c1  | 
|
2119  | 
apply(simp add: cp_def)  | 
|
2120  | 
apply(auto)  | 
|
2121  | 
apply(induct_tac x)  | 
|
2122  | 
apply(auto)  | 
|
2123  | 
done  | 
|
2124  | 
||
2125  | 
lemma cp_set_inst:  | 
|
2126  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2127  | 
  shows "cp TYPE ('a set) TYPE('x) TYPE('y)"
 | 
|
2128  | 
using c1  | 
|
2129  | 
apply(simp add: cp_def)  | 
|
2130  | 
apply(auto)  | 
|
2131  | 
apply(auto simp add: perm_set_def)  | 
|
2132  | 
apply(rule_tac x="pi2\<bullet>aa" in exI)  | 
|
2133  | 
apply(auto)  | 
|
2134  | 
done  | 
|
2135  | 
||
2136  | 
lemma cp_option_inst:  | 
|
2137  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2138  | 
  shows "cp TYPE ('a option) TYPE('x) TYPE('y)"
 | 
|
2139  | 
using c1  | 
|
2140  | 
apply(simp add: cp_def)  | 
|
2141  | 
apply(auto)  | 
|
2142  | 
apply(case_tac x)  | 
|
2143  | 
apply(auto)  | 
|
2144  | 
done  | 
|
2145  | 
||
2146  | 
lemma cp_noption_inst:  | 
|
2147  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
2148  | 
  shows "cp TYPE ('a noption) TYPE('x) TYPE('y)"
 | 
| 17870 | 2149  | 
using c1  | 
2150  | 
apply(simp add: cp_def)  | 
|
2151  | 
apply(auto)  | 
|
2152  | 
apply(case_tac x)  | 
|
2153  | 
apply(auto)  | 
|
2154  | 
done  | 
|
2155  | 
||
2156  | 
lemma cp_unit_inst:  | 
|
2157  | 
  shows "cp TYPE (unit) TYPE('x) TYPE('y)"
 | 
|
2158  | 
apply(simp add: cp_def)  | 
|
2159  | 
done  | 
|
2160  | 
||
2161  | 
lemma cp_bool_inst:  | 
|
2162  | 
  shows "cp TYPE (bool) TYPE('x) TYPE('y)"
 | 
|
2163  | 
apply(simp add: cp_def)  | 
|
2164  | 
apply(rule allI)+  | 
|
2165  | 
apply(induct_tac x)  | 
|
2166  | 
apply(simp_all)  | 
|
2167  | 
done  | 
|
2168  | 
||
2169  | 
lemma cp_prod_inst:  | 
|
2170  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2171  | 
  and     c2: "cp TYPE ('b) TYPE('x) TYPE('y)"
 | 
|
2172  | 
  shows "cp TYPE ('a\<times>'b) TYPE('x) TYPE('y)"
 | 
|
2173  | 
using c1 c2  | 
|
2174  | 
apply(simp add: cp_def)  | 
|
2175  | 
done  | 
|
2176  | 
||
2177  | 
lemma cp_fun_inst:  | 
|
2178  | 
  assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)"
 | 
|
2179  | 
  and     c2: "cp TYPE ('b) TYPE('x) TYPE('y)"
 | 
|
2180  | 
  and     pt: "pt TYPE ('y) TYPE('x)"
 | 
|
2181  | 
  and     at: "at TYPE ('x)"
 | 
|
2182  | 
  shows "cp TYPE ('a\<Rightarrow>'b) TYPE('x) TYPE('y)"
 | 
|
2183  | 
using c1 c2  | 
|
2184  | 
apply(auto simp add: cp_def perm_fun_def expand_fun_eq)  | 
|
2185  | 
apply(simp add: perm_rev[symmetric])  | 
|
2186  | 
apply(simp add: pt_rev_pi[OF pt_list_inst[OF pt_prod_inst[OF pt, OF pt]], OF at])  | 
|
2187  | 
done  | 
|
2188  | 
||
2189  | 
||
2190  | 
section {* Abstraction function *}
 | 
|
2191  | 
(*==============================*)  | 
|
2192  | 
||
2193  | 
lemma pt_abs_fun_inst:  | 
|
2194  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2195  | 
  and     at: "at TYPE('x)"
 | 
|
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
2196  | 
  shows "pt TYPE('x\<Rightarrow>('a noption)) TYPE('x)"
 | 
| 17870 | 2197  | 
by (rule pt_fun_inst[OF at_pt_inst[OF at],OF pt_noption_inst[OF pt],OF at])  | 
2198  | 
||
2199  | 
constdefs  | 
|
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
2200  | 
  abs_fun :: "'x\<Rightarrow>'a\<Rightarrow>('x\<Rightarrow>('a noption))" ("[_]._" [100,100] 100)
 | 
| 17870 | 2201  | 
"[a].x \<equiv> (\<lambda>b. (if b=a then nSome(x) else (if b\<sharp>x then nSome([(a,b)]\<bullet>x) else nNone)))"  | 
2202  | 
||
2203  | 
lemma abs_fun_if:  | 
|
2204  | 
fixes pi :: "'x prm"  | 
|
2205  | 
and x :: "'a"  | 
|
2206  | 
and y :: "'a"  | 
|
2207  | 
and c :: "bool"  | 
|
2208  | 
shows "pi\<bullet>(if c then x else y) = (if c then (pi\<bullet>x) else (pi\<bullet>y))"  | 
|
2209  | 
by force  | 
|
2210  | 
||
2211  | 
lemma abs_fun_pi_ineq:  | 
|
2212  | 
fixes a :: "'y"  | 
|
2213  | 
and x :: "'a"  | 
|
2214  | 
and pi :: "'x prm"  | 
|
2215  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
2216  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
2217  | 
  and     at:  "at TYPE('x)"
 | 
|
2218  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
2219  | 
shows "pi\<bullet>([a].x) = [(pi\<bullet>a)].(pi\<bullet>x)"  | 
|
2220  | 
apply(simp add: abs_fun_def perm_fun_def abs_fun_if)  | 
|
2221  | 
apply(simp only: expand_fun_eq)  | 
|
2222  | 
apply(rule allI)  | 
|
2223  | 
apply(subgoal_tac "(((rev pi)\<bullet>(xa::'y)) = (a::'y)) = (xa = pi\<bullet>a)")(*A*)  | 
|
2224  | 
apply(subgoal_tac "(((rev pi)\<bullet>xa)\<sharp>x) = (xa\<sharp>(pi\<bullet>x))")(*B*)  | 
|
2225  | 
apply(subgoal_tac "pi\<bullet>([(a,(rev pi)\<bullet>xa)]\<bullet>x) = [(pi\<bullet>a,xa)]\<bullet>(pi\<bullet>x)")(*C*)  | 
|
2226  | 
apply(simp)  | 
|
2227  | 
(*C*)  | 
|
2228  | 
apply(simp add: cp1[OF cp])  | 
|
2229  | 
apply(simp add: pt_pi_rev[OF ptb, OF at])  | 
|
2230  | 
(*B*)  | 
|
2231  | 
apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp])  | 
|
2232  | 
(*A*)  | 
|
2233  | 
apply(rule iffI)  | 
|
2234  | 
apply(rule pt_bij2[OF ptb, OF at, THEN sym])  | 
|
2235  | 
apply(simp)  | 
|
2236  | 
apply(rule pt_bij2[OF ptb, OF at])  | 
|
2237  | 
apply(simp)  | 
|
2238  | 
done  | 
|
2239  | 
||
2240  | 
lemma abs_fun_pi:  | 
|
2241  | 
fixes a :: "'x"  | 
|
2242  | 
and x :: "'a"  | 
|
2243  | 
and pi :: "'x prm"  | 
|
2244  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2245  | 
  and     at: "at TYPE('x)"
 | 
|
2246  | 
shows "pi\<bullet>([a].x) = [(pi\<bullet>a)].(pi\<bullet>x)"  | 
|
2247  | 
apply(rule abs_fun_pi_ineq)  | 
|
2248  | 
apply(rule pt)  | 
|
2249  | 
apply(rule at_pt_inst)  | 
|
2250  | 
apply(rule at)+  | 
|
2251  | 
apply(rule cp_pt_inst)  | 
|
2252  | 
apply(rule pt)  | 
|
2253  | 
apply(rule at)  | 
|
2254  | 
done  | 
|
2255  | 
||
2256  | 
lemma abs_fun_eq1:  | 
|
2257  | 
fixes x :: "'a"  | 
|
2258  | 
and y :: "'a"  | 
|
2259  | 
and a :: "'x"  | 
|
2260  | 
shows "([a].x = [a].y) = (x = y)"  | 
|
2261  | 
apply(auto simp add: abs_fun_def)  | 
|
2262  | 
apply(auto simp add: expand_fun_eq)  | 
|
2263  | 
apply(drule_tac x="a" in spec)  | 
|
2264  | 
apply(simp)  | 
|
2265  | 
done  | 
|
2266  | 
||
2267  | 
lemma abs_fun_eq2:  | 
|
2268  | 
fixes x :: "'a"  | 
|
2269  | 
and y :: "'a"  | 
|
2270  | 
and a :: "'x"  | 
|
2271  | 
and b :: "'x"  | 
|
2272  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2273  | 
      and at: "at TYPE('x)"
 | 
|
2274  | 
and a1: "a\<noteq>b"  | 
|
2275  | 
and a2: "[a].x = [b].y"  | 
|
| 
18268
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2276  | 
shows "x=[(a,b)]\<bullet>y \<and> a\<sharp>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2277  | 
proof -  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2278  | 
from a2 have "\<forall>c::'x. ([a].x) c = ([b].y) c" by (force simp add: expand_fun_eq)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2279  | 
hence "([a].x) a = ([b].y) a" by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2280  | 
hence a3: "nSome(x) = ([b].y) a" by (simp add: abs_fun_def)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2281  | 
show "x=[(a,b)]\<bullet>y \<and> a\<sharp>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2282  | 
proof (cases "a\<sharp>y")  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2283  | 
assume a4: "a\<sharp>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2284  | 
hence "x=[(b,a)]\<bullet>y" using a3 a1 by (simp add: abs_fun_def)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2285  | 
moreover  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2286  | 
have "[(a,b)]\<bullet>y = [(b,a)]\<bullet>y" by (rule pt3[OF pt], rule at_ds5[OF at])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2287  | 
ultimately show ?thesis using a4 by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2288  | 
next  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2289  | 
assume "\<not>a\<sharp>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2290  | 
hence "nSome(x) = nNone" using a1 a3 by (simp add: abs_fun_def)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2291  | 
hence False by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2292  | 
thus ?thesis by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2293  | 
qed  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2294  | 
qed  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2295  | 
|
| 17870 | 2296  | 
lemma abs_fun_eq3:  | 
2297  | 
fixes x :: "'a"  | 
|
2298  | 
and y :: "'a"  | 
|
2299  | 
and a :: "'x"  | 
|
2300  | 
and b :: "'x"  | 
|
2301  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2302  | 
      and at: "at TYPE('x)"
 | 
|
2303  | 
and a1: "a\<noteq>b"  | 
|
2304  | 
and a2: "x=[(a,b)]\<bullet>y"  | 
|
2305  | 
and a3: "a\<sharp>y"  | 
|
2306  | 
shows "[a].x =[b].y"  | 
|
2307  | 
proof -  | 
|
| 
18268
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2308  | 
show ?thesis  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2309  | 
proof (simp only: abs_fun_def expand_fun_eq, intro strip)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2310  | 
fix c::"'x"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2311  | 
let ?LHS = "if c=a then nSome(x) else if c\<sharp>x then nSome([(a,c)]\<bullet>x) else nNone"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2312  | 
and ?RHS = "if c=b then nSome(y) else if c\<sharp>y then nSome([(b,c)]\<bullet>y) else nNone"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2313  | 
show "?LHS=?RHS"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2314  | 
proof -  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2315  | 
have "(c=a) \<or> (c=b) \<or> (c\<noteq>a \<and> c\<noteq>b)" by blast  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2316  | 
moreover --"case c=a"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2317  | 
      { have "nSome(x) = nSome([(a,b)]\<bullet>y)" using a2 by simp
 | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2318  | 
also have "\<dots> = nSome([(b,a)]\<bullet>y)" by (simp, rule pt3[OF pt], rule at_ds5[OF at])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2319  | 
finally have "nSome(x) = nSome([(b,a)]\<bullet>y)" by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2320  | 
moreover  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2321  | 
assume "c=a"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2322  | 
ultimately have "?LHS=?RHS" using a1 a3 by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2323  | 
}  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2324  | 
moreover -- "case c=b"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2325  | 
      { have a4: "y=[(a,b)]\<bullet>x" using a2 by (simp only: pt_swap_bij[OF pt, OF at])
 | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2326  | 
hence "a\<sharp>([(a,b)]\<bullet>x)" using a3 by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2327  | 
hence "b\<sharp>x" by (simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2328  | 
moreover  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2329  | 
assume "c=b"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2330  | 
ultimately have "?LHS=?RHS" using a1 a4 by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2331  | 
}  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2332  | 
moreover -- "case c\<noteq>a \<and> c\<noteq>b"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2333  | 
      { assume a5: "c\<noteq>a \<and> c\<noteq>b"
 | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2334  | 
moreover  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2335  | 
have "c\<sharp>x = c\<sharp>y" using a2 a5 by (force simp add: at_calc[OF at] pt_fresh_left[OF pt, OF at])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2336  | 
moreover  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2337  | 
have "c\<sharp>y \<longrightarrow> [(a,c)]\<bullet>x = [(b,c)]\<bullet>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2338  | 
proof (intro strip)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2339  | 
assume a6: "c\<sharp>y"  | 
| 
18295
 
dd50de393330
changed \<sim> of permutation equality to \<triangleq>
 
urbanc 
parents: 
18294 
diff
changeset
 | 
2340  | 
have "[(a,c),(b,c),(a,c)] \<triangleq> [(a,b)]" using a1 a5 by (force intro: at_ds3[OF at])  | 
| 
18268
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2341  | 
hence "[(a,c)]\<bullet>([(b,c)]\<bullet>([(a,c)]\<bullet>y)) = [(a,b)]\<bullet>y"  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2342  | 
by (simp add: pt2[OF pt, symmetric] pt3[OF pt])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2343  | 
hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = [(a,b)]\<bullet>y" using a3 a6  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2344  | 
by (simp add: pt_fresh_fresh[OF pt, OF at])  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2345  | 
hence "[(a,c)]\<bullet>([(b,c)]\<bullet>y) = x" using a2 by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2346  | 
hence "[(b,c)]\<bullet>y = [(a,c)]\<bullet>x" by (drule_tac pt_bij1[OF pt, OF at], simp)  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2347  | 
thus "[(a,c)]\<bullet>x = [(b,c)]\<bullet>y" by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2348  | 
qed  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2349  | 
ultimately have "?LHS=?RHS" by simp  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2350  | 
}  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2351  | 
ultimately show "?LHS = ?RHS" by blast  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2352  | 
qed  | 
| 17870 | 2353  | 
qed  | 
| 
18268
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2354  | 
qed  | 
| 
 
734f23ad5d8f
ISAR-fied two proofs about equality for abstraction functions.
 
urbanc 
parents: 
18264 
diff
changeset
 | 
2355  | 
|
| 17870 | 2356  | 
lemma abs_fun_eq:  | 
2357  | 
fixes x :: "'a"  | 
|
2358  | 
and y :: "'a"  | 
|
2359  | 
and a :: "'x"  | 
|
2360  | 
and b :: "'x"  | 
|
2361  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2362  | 
      and at: "at TYPE('x)"
 | 
|
2363  | 
shows "([a].x = [b].y) = ((a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y))"  | 
|
2364  | 
proof (rule iffI)  | 
|
2365  | 
assume b: "[a].x = [b].y"  | 
|
2366  | 
show "(a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y)"  | 
|
2367  | 
proof (cases "a=b")  | 
|
2368  | 
case True with b show ?thesis by (simp add: abs_fun_eq1)  | 
|
2369  | 
next  | 
|
2370  | 
case False with b show ?thesis by (simp add: abs_fun_eq2[OF pt, OF at])  | 
|
2371  | 
qed  | 
|
2372  | 
next  | 
|
2373  | 
assume "(a=b \<and> x=y)\<or>(a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y)"  | 
|
2374  | 
thus "[a].x = [b].y"  | 
|
2375  | 
proof  | 
|
2376  | 
assume "a=b \<and> x=y" thus ?thesis by simp  | 
|
2377  | 
next  | 
|
2378  | 
assume "a\<noteq>b \<and> x=[(a,b)]\<bullet>y \<and> a\<sharp>y"  | 
|
2379  | 
thus ?thesis by (simp add: abs_fun_eq3[OF pt, OF at])  | 
|
2380  | 
qed  | 
|
2381  | 
qed  | 
|
2382  | 
||
2383  | 
lemma abs_fun_supp_approx:  | 
|
2384  | 
fixes x :: "'a"  | 
|
2385  | 
and a :: "'x"  | 
|
2386  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2387  | 
  and     at: "at TYPE('x)"
 | 
|
| 18048 | 2388  | 
shows "((supp ([a].x))::'x set) \<subseteq> (supp (x,a))"  | 
2389  | 
proof  | 
|
2390  | 
fix c  | 
|
2391  | 
assume "c\<in>((supp ([a].x))::'x set)"  | 
|
2392  | 
  hence "infinite {b. [(c,b)]\<bullet>([a].x) \<noteq> [a].x}" by (simp add: supp_def)
 | 
|
2393  | 
  hence "infinite {b. [([(c,b)]\<bullet>a)].([(c,b)]\<bullet>x) \<noteq> [a].x}" by (simp add: abs_fun_pi[OF pt, OF at])
 | 
|
2394  | 
moreover  | 
|
2395  | 
  have "{b. [([(c,b)]\<bullet>a)].([(c,b)]\<bullet>x) \<noteq> [a].x} \<subseteq> {b. ([(c,b)]\<bullet>x,[(c,b)]\<bullet>a) \<noteq> (x, a)}" by force
 | 
|
2396  | 
  ultimately have "infinite {b. ([(c,b)]\<bullet>x,[(c,b)]\<bullet>a) \<noteq> (x, a)}" by (simp add: infinite_super)
 | 
|
2397  | 
thus "c\<in>(supp (x,a))" by (simp add: supp_def)  | 
|
| 17870 | 2398  | 
qed  | 
2399  | 
||
2400  | 
lemma abs_fun_finite_supp:  | 
|
2401  | 
fixes x :: "'a"  | 
|
2402  | 
and a :: "'x"  | 
|
2403  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2404  | 
  and     at: "at TYPE('x)"
 | 
|
2405  | 
and f: "finite ((supp x)::'x set)"  | 
|
2406  | 
shows "finite ((supp ([a].x))::'x set)"  | 
|
2407  | 
proof -  | 
|
| 18048 | 2408  | 
from f have "finite ((supp (x,a))::'x set)" by (simp add: supp_prod at_supp[OF at])  | 
2409  | 
moreover  | 
|
2410  | 
have "((supp ([a].x))::'x set) \<subseteq> (supp (x,a))" by (rule abs_fun_supp_approx[OF pt, OF at])  | 
|
2411  | 
ultimately show ?thesis by (simp add: finite_subset)  | 
|
| 17870 | 2412  | 
qed  | 
2413  | 
||
2414  | 
lemma fresh_abs_funI1:  | 
|
2415  | 
fixes x :: "'a"  | 
|
2416  | 
and a :: "'x"  | 
|
2417  | 
and b :: "'x"  | 
|
2418  | 
  assumes pt:  "pt TYPE('a) TYPE('x)"
 | 
|
2419  | 
  and     at:   "at TYPE('x)"
 | 
|
2420  | 
and f: "finite ((supp x)::'x set)"  | 
|
2421  | 
and a1: "b\<sharp>x"  | 
|
2422  | 
and a2: "a\<noteq>b"  | 
|
2423  | 
shows "b\<sharp>([a].x)"  | 
|
2424  | 
proof -  | 
|
2425  | 
have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)"  | 
|
2426  | 
proof (rule at_exists_fresh[OF at], auto simp add: supp_prod at_supp[OF at] f)  | 
|
2427  | 
show "finite ((supp ([a].x))::'x set)" using f  | 
|
2428  | 
by (simp add: abs_fun_finite_supp[OF pt, OF at])  | 
|
2429  | 
qed  | 
|
2430  | 
then obtain c where fr1: "c\<noteq>b"  | 
|
2431  | 
and fr2: "c\<noteq>a"  | 
|
2432  | 
and fr3: "c\<sharp>x"  | 
|
2433  | 
and fr4: "c\<sharp>([a].x)"  | 
|
2434  | 
by (force simp add: fresh_prod at_fresh[OF at])  | 
|
2435  | 
have e: "[(c,b)]\<bullet>([a].x) = [a].([(c,b)]\<bullet>x)" using a2 fr1 fr2  | 
|
2436  | 
by (force simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])  | 
|
2437  | 
from fr4 have "([(c,b)]\<bullet>c)\<sharp> ([(c,b)]\<bullet>([a].x))"  | 
|
2438  | 
by (simp add: pt_fresh_bij[OF pt_abs_fun_inst[OF pt, OF at], OF at])  | 
|
2439  | 
hence "b\<sharp>([a].([(c,b)]\<bullet>x))" using fr1 fr2 e  | 
|
2440  | 
by (simp add: at_calc[OF at])  | 
|
2441  | 
thus ?thesis using a1 fr3  | 
|
2442  | 
by (simp add: pt_fresh_fresh[OF pt, OF at])  | 
|
2443  | 
qed  | 
|
2444  | 
||
2445  | 
lemma fresh_abs_funE:  | 
|
2446  | 
fixes a :: "'x"  | 
|
2447  | 
and b :: "'x"  | 
|
2448  | 
and x :: "'a"  | 
|
2449  | 
  assumes pt:  "pt TYPE('a) TYPE('x)"
 | 
|
2450  | 
  and     at:  "at TYPE('x)"
 | 
|
2451  | 
and f: "finite ((supp x)::'x set)"  | 
|
2452  | 
and a1: "b\<sharp>([a].x)"  | 
|
2453  | 
and a2: "b\<noteq>a"  | 
|
2454  | 
shows "b\<sharp>x"  | 
|
2455  | 
proof -  | 
|
2456  | 
have "\<exists>c::'x. c\<sharp>(b,a,x,[a].x)"  | 
|
2457  | 
proof (rule at_exists_fresh[OF at], auto simp add: supp_prod at_supp[OF at] f)  | 
|
2458  | 
show "finite ((supp ([a].x))::'x set)" using f  | 
|
2459  | 
by (simp add: abs_fun_finite_supp[OF pt, OF at])  | 
|
2460  | 
qed  | 
|
2461  | 
then obtain c where fr1: "b\<noteq>c"  | 
|
2462  | 
and fr2: "c\<noteq>a"  | 
|
2463  | 
and fr3: "c\<sharp>x"  | 
|
2464  | 
and fr4: "c\<sharp>([a].x)" by (force simp add: fresh_prod at_fresh[OF at])  | 
|
2465  | 
have "[a].x = [(b,c)]\<bullet>([a].x)" using a1 fr4  | 
|
2466  | 
by (simp add: pt_fresh_fresh[OF pt_abs_fun_inst[OF pt, OF at], OF at])  | 
|
2467  | 
hence "[a].x = [a].([(b,c)]\<bullet>x)" using fr2 a2  | 
|
2468  | 
by (force simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])  | 
|
2469  | 
hence b: "([(b,c)]\<bullet>x) = x" by (simp add: abs_fun_eq1)  | 
|
2470  | 
from fr3 have "([(b,c)]\<bullet>c)\<sharp>([(b,c)]\<bullet>x)"  | 
|
2471  | 
by (simp add: pt_fresh_bij[OF pt, OF at])  | 
|
2472  | 
thus ?thesis using b fr1 by (simp add: at_calc[OF at])  | 
|
2473  | 
qed  | 
|
2474  | 
||
2475  | 
lemma fresh_abs_funI2:  | 
|
2476  | 
fixes a :: "'x"  | 
|
2477  | 
and x :: "'a"  | 
|
2478  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2479  | 
  and     at: "at TYPE('x)"
 | 
|
2480  | 
and f: "finite ((supp x)::'x set)"  | 
|
2481  | 
shows "a\<sharp>([a].x)"  | 
|
2482  | 
proof -  | 
|
2483  | 
have "\<exists>c::'x. c\<sharp>(a,x)"  | 
|
2484  | 
by (rule at_exists_fresh[OF at], auto simp add: supp_prod at_supp[OF at] f)  | 
|
2485  | 
then obtain c where fr1: "a\<noteq>c" and fr1_sym: "c\<noteq>a"  | 
|
2486  | 
and fr2: "c\<sharp>x" by (force simp add: fresh_prod at_fresh[OF at])  | 
|
2487  | 
have "c\<sharp>([a].x)" using f fr1 fr2 by (simp add: fresh_abs_funI1[OF pt, OF at])  | 
|
2488  | 
hence "([(c,a)]\<bullet>c)\<sharp>([(c,a)]\<bullet>([a].x))" using fr1  | 
|
2489  | 
by (simp only: pt_fresh_bij[OF pt_abs_fun_inst[OF pt, OF at], OF at])  | 
|
2490  | 
hence a: "a\<sharp>([c].([(c,a)]\<bullet>x))" using fr1_sym  | 
|
2491  | 
by (simp add: abs_fun_pi[OF pt, OF at] at_calc[OF at])  | 
|
2492  | 
have "[c].([(c,a)]\<bullet>x) = ([a].x)" using fr1_sym fr2  | 
|
2493  | 
by (simp add: abs_fun_eq[OF pt, OF at])  | 
|
2494  | 
thus ?thesis using a by simp  | 
|
2495  | 
qed  | 
|
2496  | 
||
2497  | 
lemma fresh_abs_fun_iff:  | 
|
2498  | 
fixes a :: "'x"  | 
|
2499  | 
and b :: "'x"  | 
|
2500  | 
and x :: "'a"  | 
|
2501  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2502  | 
  and     at: "at TYPE('x)"
 | 
|
2503  | 
and f: "finite ((supp x)::'x set)"  | 
|
2504  | 
shows "(b\<sharp>([a].x)) = (b=a \<or> b\<sharp>x)"  | 
|
2505  | 
by (auto dest: fresh_abs_funE[OF pt, OF at,OF f]  | 
|
2506  | 
intro: fresh_abs_funI1[OF pt, OF at,OF f]  | 
|
2507  | 
fresh_abs_funI2[OF pt, OF at,OF f])  | 
|
2508  | 
||
2509  | 
lemma abs_fun_supp:  | 
|
2510  | 
fixes a :: "'x"  | 
|
2511  | 
and x :: "'a"  | 
|
2512  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2513  | 
  and     at: "at TYPE('x)"
 | 
|
2514  | 
and f: "finite ((supp x)::'x set)"  | 
|
2515  | 
  shows "supp ([a].x) = (supp x)-{a}"
 | 
|
2516  | 
by (force simp add: supp_fresh_iff fresh_abs_fun_iff[OF pt, OF at, OF f])  | 
|
2517  | 
||
| 18048 | 2518  | 
(* maybe needs to be better stated as supp intersection supp *)  | 
| 17870 | 2519  | 
lemma abs_fun_supp_ineq:  | 
2520  | 
fixes a :: "'y"  | 
|
2521  | 
and x :: "'a"  | 
|
2522  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
2523  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
2524  | 
  and     at:  "at TYPE('x)"
 | 
|
2525  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
2526  | 
  and     dj:  "disjoint TYPE('y) TYPE('x)"
 | 
|
2527  | 
shows "((supp ([a].x))::'x set) = (supp x)"  | 
|
2528  | 
apply(auto simp add: supp_def)  | 
|
2529  | 
apply(auto simp add: abs_fun_pi_ineq[OF pta, OF ptb, OF at, OF cp])  | 
|
2530  | 
apply(auto simp add: dj_perm_forget[OF dj])  | 
|
2531  | 
apply(auto simp add: abs_fun_eq1)  | 
|
2532  | 
done  | 
|
2533  | 
||
2534  | 
lemma fresh_abs_fun_iff_ineq:  | 
|
2535  | 
fixes a :: "'y"  | 
|
2536  | 
and b :: "'x"  | 
|
2537  | 
and x :: "'a"  | 
|
2538  | 
  assumes pta: "pt TYPE('a) TYPE('x)"
 | 
|
2539  | 
  and     ptb: "pt TYPE('y) TYPE('x)"
 | 
|
2540  | 
  and     at:  "at TYPE('x)"
 | 
|
2541  | 
  and     cp:  "cp TYPE('a) TYPE('x) TYPE('y)"
 | 
|
2542  | 
  and     dj:  "disjoint TYPE('y) TYPE('x)"
 | 
|
2543  | 
shows "b\<sharp>([a].x) = b\<sharp>x"  | 
|
2544  | 
by (simp add: fresh_def abs_fun_supp_ineq[OF pta, OF ptb, OF at, OF cp, OF dj])  | 
|
2545  | 
||
| 18048 | 2546  | 
section {* abstraction type for the parsing in nominal datatype *}
 | 
2547  | 
(*==============================================================*)  | 
|
| 17870 | 2548  | 
consts  | 
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
2549  | 
  "ABS_set" :: "('x\<Rightarrow>('a noption)) set"
 | 
| 17870 | 2550  | 
inductive ABS_set  | 
2551  | 
intros  | 
|
2552  | 
ABS_in: "(abs_fun a x)\<in>ABS_set"  | 
|
2553  | 
||
| 
18579
 
002d371401f5
changed the name of the type "nOption" to "noption".
 
urbanc 
parents: 
18578 
diff
changeset
 | 
2554  | 
typedef (ABS) ('x,'a) ABS = "ABS_set::('x\<Rightarrow>('a noption)) set"
 | 
| 17870 | 2555  | 
proof  | 
2556  | 
fix x::"'a" and a::"'x"  | 
|
2557  | 
show "(abs_fun a x)\<in> ABS_set" by (rule ABS_in)  | 
|
2558  | 
qed  | 
|
2559  | 
||
2560  | 
syntax ABS :: "type \<Rightarrow> type \<Rightarrow> type" ("\<guillemotleft>_\<guillemotright>_" [1000,1000] 1000)
 | 
|
2561  | 
||
2562  | 
||
| 18048 | 2563  | 
section {* lemmas for deciding permutation equations *}
 | 
| 17870 | 2564  | 
(*===================================================*)  | 
2565  | 
||
2566  | 
lemma perm_eq_app:  | 
|
2567  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
2568  | 
and x :: "'a"  | 
|
2569  | 
and pi :: "'x prm"  | 
|
2570  | 
  assumes pt: "pt TYPE('a) TYPE('x)"
 | 
|
2571  | 
  and     at: "at TYPE('x)"
 | 
|
2572  | 
shows "(pi\<bullet>(f x)=y) = ((pi\<bullet>f)(pi\<bullet>x)=y)"  | 
|
2573  | 
by (simp add: pt_fun_app_eq[OF pt, OF at])  | 
|
2574  | 
||
2575  | 
lemma perm_eq_lam:  | 
|
2576  | 
fixes f :: "'a\<Rightarrow>'b"  | 
|
2577  | 
and x :: "'a"  | 
|
2578  | 
and pi :: "'x prm"  | 
|
2579  | 
shows "((pi\<bullet>(\<lambda>x. f x))=y) = ((\<lambda>x. (pi\<bullet>(f ((rev pi)\<bullet>x))))=y)"  | 
|
2580  | 
by (simp add: perm_fun_def)  | 
|
2581  | 
||
2582  | 
||
2583  | 
(***************************************)  | 
|
2584  | 
(* setup for the individial atom-kinds *)  | 
|
| 
18047
 
3d643b13eb65
simplified the abs_supp_approx proof and tuned some comments in
 
urbanc 
parents: 
18012 
diff
changeset
 | 
2585  | 
(* and nominal datatypes *)  | 
| 18068 | 2586  | 
use "nominal_atoms.ML"  | 
| 17870 | 2587  | 
use "nominal_package.ML"  | 
| 18068 | 2588  | 
setup "NominalAtoms.setup"  | 
| 17870 | 2589  | 
|
| 
18047
 
3d643b13eb65
simplified the abs_supp_approx proof and tuned some comments in
 
urbanc 
parents: 
18012 
diff
changeset
 | 
2590  | 
(*****************************************)  | 
| 
 
3d643b13eb65
simplified the abs_supp_approx proof and tuned some comments in
 
urbanc 
parents: 
18012 
diff
changeset
 | 
2591  | 
(* setup for induction principles method *)  | 
| 
18294
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
2592  | 
|
| 17870 | 2593  | 
use "nominal_induct.ML";  | 
2594  | 
method_setup nominal_induct =  | 
|
| 
18294
 
bbfd64cc91ab
fresh_unit_elim and fresh_prod_elim -- for nominal_induct;
 
wenzelm 
parents: 
18268 
diff
changeset
 | 
2595  | 
  {* NominalInduct.nominal_induct_method *}
 | 
| 17870 | 2596  | 
  {* nominal induction *}
 | 
2597  | 
||
2598  | 
(*******************************)  | 
|
2599  | 
(* permutation equality tactic *)  | 
|
2600  | 
use "nominal_permeq.ML";  | 
|
| 18012 | 2601  | 
|
| 17870 | 2602  | 
method_setup perm_simp =  | 
2603  | 
  {* perm_eq_meth *}
 | 
|
2604  | 
  {* tactic for deciding equalities involving permutations *}
 | 
|
2605  | 
||
2606  | 
method_setup perm_simp_debug =  | 
|
2607  | 
  {* perm_eq_meth_debug *}
 | 
|
| 
18047
 
3d643b13eb65
simplified the abs_supp_approx proof and tuned some comments in
 
urbanc 
parents: 
18012 
diff
changeset
 | 
2608  | 
  {* tactic for deciding equalities involving permutations including debuging facilities *}
 | 
| 17870 | 2609  | 
|
2610  | 
method_setup supports_simp =  | 
|
2611  | 
  {* supports_meth *}
 | 
|
2612  | 
  {* tactic for deciding whether something supports semthing else *}
 | 
|
2613  | 
||
2614  | 
method_setup supports_simp_debug =  | 
|
2615  | 
  {* supports_meth_debug *}
 | 
|
| 
18047
 
3d643b13eb65
simplified the abs_supp_approx proof and tuned some comments in
 
urbanc 
parents: 
18012 
diff
changeset
 | 
2616  | 
  {* tactic for deciding equalities involving permutations including debuging facilities *}
 | 
| 17870 | 2617  | 
|
2618  | 
end  |