|
1 theory LamEx |
|
2 imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "../QuotMain" |
|
3 begin |
|
4 |
|
5 |
|
6 (* lemmas that should be in Nominal \<dots>\<dots>must be cleaned *) |
|
7 lemma in_permute_iff: |
|
8 shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X" |
|
9 apply(unfold mem_def permute_fun_def)[1] |
|
10 apply(simp add: permute_bool_def) |
|
11 done |
|
12 |
|
13 lemma fresh_star_permute_iff: |
|
14 shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x" |
|
15 apply(simp add: fresh_star_def) |
|
16 apply(auto) |
|
17 apply(drule_tac x="p \<bullet> xa" in bspec) |
|
18 apply(unfold mem_def permute_fun_def)[1] |
|
19 apply(simp add: eqvts) |
|
20 apply(simp add: fresh_permute_iff) |
|
21 apply(rule_tac ?p1="- p" in fresh_permute_iff[THEN iffD1]) |
|
22 apply(simp) |
|
23 apply(drule_tac x="- p \<bullet> xa" in bspec) |
|
24 apply(rule_tac ?p1="p" in in_permute_iff[THEN iffD1]) |
|
25 apply(simp) |
|
26 apply(simp) |
|
27 done |
|
28 |
|
29 lemma fresh_minus_perm: |
|
30 fixes p::perm |
|
31 shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p" |
|
32 apply(simp add: fresh_def) |
|
33 apply(simp only: supp_minus_perm) |
|
34 done |
|
35 |
|
36 lemma fresh_plus: |
|
37 fixes p q::perm |
|
38 shows "\<lbrakk>a \<sharp> p; a \<sharp> q\<rbrakk> \<Longrightarrow> a \<sharp> (p + q)" |
|
39 unfolding fresh_def |
|
40 using supp_plus_perm |
|
41 apply(auto) |
|
42 done |
|
43 |
|
44 lemma fresh_star_plus: |
|
45 fixes p q::perm |
|
46 shows "\<lbrakk>a \<sharp>* p; a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)" |
|
47 unfolding fresh_star_def |
|
48 by (simp add: fresh_plus) |
|
49 |
|
50 |
|
51 |
|
52 atom_decl name |
|
53 |
|
54 datatype rlam = |
|
55 rVar "name" |
|
56 | rApp "rlam" "rlam" |
|
57 | rLam "name" "rlam" |
|
58 |
|
59 fun |
|
60 rfv :: "rlam \<Rightarrow> atom set" |
|
61 where |
|
62 rfv_var: "rfv (rVar a) = {atom a}" |
|
63 | rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)" |
|
64 | rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}" |
|
65 |
|
66 instantiation rlam :: pt |
|
67 begin |
|
68 |
|
69 primrec |
|
70 permute_rlam |
|
71 where |
|
72 "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)" |
|
73 | "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)" |
|
74 | "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)" |
|
75 |
|
76 instance |
|
77 apply default |
|
78 apply(induct_tac [!] x) |
|
79 apply(simp_all) |
|
80 done |
|
81 |
|
82 end |
|
83 |
|
84 instantiation rlam :: fs |
|
85 begin |
|
86 |
|
87 lemma neg_conj: |
|
88 "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)" |
|
89 by simp |
|
90 |
|
91 lemma infinite_Un: |
|
92 "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T" |
|
93 by simp |
|
94 |
|
95 instance |
|
96 apply default |
|
97 apply(induct_tac x) |
|
98 (* var case *) |
|
99 apply(simp add: supp_def) |
|
100 apply(fold supp_def)[1] |
|
101 apply(simp add: supp_at_base) |
|
102 (* app case *) |
|
103 apply(simp only: supp_def) |
|
104 apply(simp only: permute_rlam.simps) |
|
105 apply(simp only: rlam.inject) |
|
106 apply(simp only: neg_conj) |
|
107 apply(simp only: Collect_disj_eq) |
|
108 apply(simp only: infinite_Un) |
|
109 apply(simp only: Collect_disj_eq) |
|
110 apply(simp) |
|
111 (* lam case *) |
|
112 apply(simp only: supp_def) |
|
113 apply(simp only: permute_rlam.simps) |
|
114 apply(simp only: rlam.inject) |
|
115 apply(simp only: neg_conj) |
|
116 apply(simp only: Collect_disj_eq) |
|
117 apply(simp only: infinite_Un) |
|
118 apply(simp only: Collect_disj_eq) |
|
119 apply(simp) |
|
120 apply(fold supp_def)[1] |
|
121 apply(simp add: supp_at_base) |
|
122 done |
|
123 |
|
124 end |
|
125 |
|
126 |
|
127 (* for the eqvt proof of the alpha-equivalence *) |
|
128 declare permute_rlam.simps[eqvt] |
|
129 |
|
130 lemma rfv_eqvt[eqvt]: |
|
131 shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)" |
|
132 apply(induct t) |
|
133 apply(simp_all) |
|
134 apply(simp add: permute_set_eq atom_eqvt) |
|
135 apply(simp add: union_eqvt) |
|
136 apply(simp add: Diff_eqvt) |
|
137 apply(simp add: permute_set_eq atom_eqvt) |
|
138 done |
|
139 |
|
140 inductive |
|
141 alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100) |
|
142 where |
|
143 a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)" |
|
144 | a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" |
|
145 | a3: "\<exists>pi. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b) |
|
146 \<Longrightarrow> rLam a t \<approx> rLam b s" |
|
147 |
|
148 text {* should be automatic with new version of eqvt-machinery *} |
|
149 lemma alpha_eqvt: |
|
150 shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)" |
|
151 apply(induct rule: alpha.induct) |
|
152 apply(simp add: a1) |
|
153 apply(simp add: a2) |
|
154 apply(simp) |
|
155 apply(rule a3) |
|
156 apply(erule conjE) |
|
157 apply(erule exE) |
|
158 apply(erule conjE) |
|
159 apply(rule_tac x="pi \<bullet> pia" in exI) |
|
160 apply(rule conjI) |
|
161 apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) |
|
162 apply(simp add: eqvts atom_eqvt) |
|
163 apply(rule conjI) |
|
164 apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) |
|
165 apply(simp add: eqvts atom_eqvt) |
|
166 apply(rule conjI) |
|
167 apply(subst permute_eqvt[symmetric]) |
|
168 apply(simp) |
|
169 apply(subst permute_eqvt[symmetric]) |
|
170 apply(simp) |
|
171 done |
|
172 |
|
173 lemma alpha_refl: |
|
174 shows "t \<approx> t" |
|
175 apply(induct t rule: rlam.induct) |
|
176 apply(simp add: a1) |
|
177 apply(simp add: a2) |
|
178 apply(rule a3) |
|
179 apply(rule_tac x="0" in exI) |
|
180 apply(simp_all add: fresh_star_def fresh_zero_perm) |
|
181 done |
|
182 |
|
183 lemma alpha_sym: |
|
184 shows "t \<approx> s \<Longrightarrow> s \<approx> t" |
|
185 apply(induct rule: alpha.induct) |
|
186 apply(simp add: a1) |
|
187 apply(simp add: a2) |
|
188 apply(rule a3) |
|
189 apply(erule exE) |
|
190 apply(rule_tac x="- pi" in exI) |
|
191 apply(simp) |
|
192 apply(simp add: fresh_star_def fresh_minus_perm) |
|
193 apply(rule conjI) |
|
194 apply(erule conjE)+ |
|
195 apply(rotate_tac 3) |
|
196 apply(drule_tac pi="- pi" in alpha_eqvt) |
|
197 apply(simp) |
|
198 apply(clarify) |
|
199 apply(simp) |
|
200 done |
|
201 |
|
202 lemma alpha_trans: |
|
203 shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3" |
|
204 apply(induct arbitrary: t3 rule: alpha.induct) |
|
205 apply(erule alpha.cases) |
|
206 apply(simp_all) |
|
207 apply(simp add: a1) |
|
208 apply(rotate_tac 4) |
|
209 apply(erule alpha.cases) |
|
210 apply(simp_all) |
|
211 apply(simp add: a2) |
|
212 apply(rotate_tac 1) |
|
213 apply(erule alpha.cases) |
|
214 apply(simp_all) |
|
215 apply(erule conjE)+ |
|
216 apply(erule exE)+ |
|
217 apply(erule conjE)+ |
|
218 apply(rule a3) |
|
219 apply(rule_tac x="pia + pi" in exI) |
|
220 apply(simp add: fresh_star_plus) |
|
221 apply(drule_tac x="- pia \<bullet> sa" in spec) |
|
222 apply(drule mp) |
|
223 apply(rotate_tac 8) |
|
224 apply(drule_tac pi="- pia" in alpha_eqvt) |
|
225 apply(simp) |
|
226 apply(rotate_tac 11) |
|
227 apply(drule_tac pi="pia" in alpha_eqvt) |
|
228 apply(simp) |
|
229 done |
|
230 |
|
231 lemma alpha_equivp: |
|
232 shows "equivp alpha" |
|
233 apply(rule equivpI) |
|
234 unfolding reflp_def symp_def transp_def |
|
235 apply(auto intro: alpha_refl alpha_sym alpha_trans) |
|
236 done |
|
237 |
|
238 lemma alpha_rfv: |
|
239 shows "t \<approx> s \<Longrightarrow> rfv t = rfv s" |
|
240 apply(induct rule: alpha.induct) |
|
241 apply(simp) |
|
242 apply(simp) |
|
243 apply(simp) |
|
244 done |
|
245 |
|
246 (* PROBABLY NOT TRUE !!! needed for lifting |
|
247 lemma alpha_fresh: |
|
248 assumes a: "t \<approx> s" |
|
249 shows "a\<sharp>t \<Longrightarrow> a\<sharp>s" |
|
250 using a |
|
251 apply(induct) |
|
252 *) |
|
253 |
|
254 quotient_type lam = rlam / alpha |
|
255 by (rule alpha_equivp) |
|
256 |
|
257 quotient_definition |
|
258 "Var :: name \<Rightarrow> lam" |
|
259 as |
|
260 "rVar" |
|
261 |
|
262 quotient_definition |
|
263 "App :: lam \<Rightarrow> lam \<Rightarrow> lam" |
|
264 as |
|
265 "rApp" |
|
266 |
|
267 quotient_definition |
|
268 "Lam :: name \<Rightarrow> lam \<Rightarrow> lam" |
|
269 as |
|
270 "rLam" |
|
271 |
|
272 quotient_definition |
|
273 "fv :: lam \<Rightarrow> atom set" |
|
274 as |
|
275 "rfv" |
|
276 |
|
277 lemma perm_rsp[quot_respect]: |
|
278 "(op = ===> alpha ===> alpha) permute permute" |
|
279 apply(auto) |
|
280 apply(rule alpha_eqvt) |
|
281 apply(simp) |
|
282 done |
|
283 |
|
284 lemma rVar_rsp[quot_respect]: |
|
285 "(op = ===> alpha) rVar rVar" |
|
286 by (auto intro: a1) |
|
287 |
|
288 lemma rApp_rsp[quot_respect]: |
|
289 "(alpha ===> alpha ===> alpha) rApp rApp" |
|
290 by (auto intro: a2) |
|
291 |
|
292 lemma rLam_rsp[quot_respect]: |
|
293 "(op = ===> alpha ===> alpha) rLam rLam" |
|
294 apply(auto) |
|
295 apply(rule a3) |
|
296 apply(rule_tac x="0" in exI) |
|
297 unfolding fresh_star_def |
|
298 apply(simp add: fresh_star_def fresh_zero_perm) |
|
299 apply(simp add: alpha_rfv) |
|
300 done |
|
301 |
|
302 lemma rfv_rsp[quot_respect]: |
|
303 "(alpha ===> op =) rfv rfv" |
|
304 apply(simp add: alpha_rfv) |
|
305 done |
|
306 |
|
307 |
|
308 |
|
309 (* NOT SURE see above |
|
310 lemma fresh_rsp: |
|
311 "(op = ===> alpha ===> op =) fresh fresh" |
|
312 apply(auto) |
|
313 *) |
|
314 |
|
315 section {* lifted theorems *} |
|
316 |
|
317 lemma lam_induct: |
|
318 "\<lbrakk>\<And>name. P (Var name); |
|
319 \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2); |
|
320 \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> |
|
321 \<Longrightarrow> P lam" |
|
322 apply (lifting rlam.induct) |
|
323 done |
|
324 |
|
325 instantiation lam :: pt |
|
326 begin |
|
327 |
|
328 quotient_definition |
|
329 "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam" |
|
330 as |
|
331 "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam" |
|
332 |
|
333 lemma permute_lam [simp]: |
|
334 shows "pi \<bullet> Var a = Var (pi \<bullet> a)" |
|
335 and "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)" |
|
336 and "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)" |
|
337 apply(lifting permute_rlam.simps) |
|
338 done |
|
339 |
|
340 instance |
|
341 apply default |
|
342 apply(induct_tac [!] x rule: lam_induct) |
|
343 apply(simp_all) |
|
344 done |
|
345 |
|
346 end |
|
347 |
|
348 lemma fv_lam [simp]: |
|
349 shows "fv (Var a) = {atom a}" |
|
350 and "fv (App t1 t2) = fv t1 \<union> fv t2" |
|
351 and "fv (Lam a t) = fv t - {atom a}" |
|
352 apply(lifting rfv_var rfv_app rfv_lam) |
|
353 done |
|
354 |
|
355 lemma a1: |
|
356 "a = b \<Longrightarrow> Var a = Var b" |
|
357 by (lifting a1) |
|
358 |
|
359 lemma a2: |
|
360 "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc" |
|
361 by (lifting a2) |
|
362 |
|
363 lemma a3: |
|
364 "\<lbrakk>\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk> |
|
365 \<Longrightarrow> Lam a t = Lam b s" |
|
366 apply (lifting a3) |
|
367 done |
|
368 |
|
369 lemma alpha_cases: |
|
370 "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; |
|
371 \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; |
|
372 \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; |
|
373 \<exists>pi. fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> |
|
374 \<Longrightarrow> P\<rbrakk> |
|
375 \<Longrightarrow> P" |
|
376 by (lifting alpha.cases) |
|
377 |
|
378 lemma alpha_induct: |
|
379 "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b); |
|
380 \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc); |
|
381 \<And>t a s b. |
|
382 \<lbrakk>\<exists>pi. fv t - {atom a} = fv s - {atom b} \<and> |
|
383 (fv t - {atom a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> |
|
384 \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk> |
|
385 \<Longrightarrow> qxb qx qxa" |
|
386 by (lifting alpha.induct) |
|
387 |
|
388 (* should they lift automatically *) |
|
389 lemma lam_inject [simp]: |
|
390 shows "(Var a = Var b) = (a = b)" |
|
391 and "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)" |
|
392 apply(lifting rlam.inject(1) rlam.inject(2)) |
|
393 apply(auto) |
|
394 apply(drule alpha.cases) |
|
395 apply(simp_all) |
|
396 apply(simp add: alpha.a1) |
|
397 apply(drule alpha.cases) |
|
398 apply(simp_all) |
|
399 apply(drule alpha.cases) |
|
400 apply(simp_all) |
|
401 apply(rule alpha.a2) |
|
402 apply(simp_all) |
|
403 done |
|
404 |
|
405 lemma rlam_distinct: |
|
406 shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')" |
|
407 and "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)" |
|
408 and "\<not>(rVar nam \<approx> rLam nam' rlam')" |
|
409 and "\<not>(rLam nam' rlam' \<approx> rVar nam)" |
|
410 and "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')" |
|
411 and "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)" |
|
412 apply auto |
|
413 apply(erule alpha.cases) |
|
414 apply simp_all |
|
415 apply(erule alpha.cases) |
|
416 apply simp_all |
|
417 apply(erule alpha.cases) |
|
418 apply simp_all |
|
419 apply(erule alpha.cases) |
|
420 apply simp_all |
|
421 apply(erule alpha.cases) |
|
422 apply simp_all |
|
423 apply(erule alpha.cases) |
|
424 apply simp_all |
|
425 done |
|
426 |
|
427 lemma lam_distinct[simp]: |
|
428 shows "Var nam \<noteq> App lam1' lam2'" |
|
429 and "App lam1' lam2' \<noteq> Var nam" |
|
430 and "Var nam \<noteq> Lam nam' lam'" |
|
431 and "Lam nam' lam' \<noteq> Var nam" |
|
432 and "App lam1 lam2 \<noteq> Lam nam' lam'" |
|
433 and "Lam nam' lam' \<noteq> App lam1 lam2" |
|
434 apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6)) |
|
435 done |
|
436 |
|
437 lemma var_supp1: |
|
438 shows "(supp (Var a)) = (supp a)" |
|
439 apply (simp add: supp_def) |
|
440 done |
|
441 |
|
442 lemma var_supp: |
|
443 shows "(supp (Var a)) = {a:::name}" |
|
444 using var_supp1 by (simp add: supp_at_base) |
|
445 |
|
446 lemma app_supp: |
|
447 shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)" |
|
448 apply(simp only: permute_lam supp_def lam_inject) |
|
449 apply(simp add: Collect_imp_eq Collect_neg_eq) |
|
450 done |
|
451 |
|
452 (* needs thinking *) |
|
453 lemma lam_supp: |
|
454 shows "supp (Lam x t) = ((supp t) - {atom x})" |
|
455 apply(simp add: supp_def) |
|
456 sorry |
|
457 |
|
458 |
|
459 instance lam :: fs |
|
460 apply(default) |
|
461 apply(induct_tac x rule: lam_induct) |
|
462 apply(simp add: var_supp) |
|
463 apply(simp add: app_supp) |
|
464 apply(simp add: lam_supp) |
|
465 done |
|
466 |
|
467 lemma fresh_lam: |
|
468 "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)" |
|
469 apply(simp add: fresh_def) |
|
470 apply(simp add: lam_supp) |
|
471 apply(auto) |
|
472 done |
|
473 |
|
474 lemma lam_induct_strong: |
|
475 fixes a::"'a::fs" |
|
476 assumes a1: "\<And>name b. P b (Var name)" |
|
477 and a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)" |
|
478 and a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)" |
|
479 shows "P a lam" |
|
480 proof - |
|
481 have "\<And>pi a. P a (pi \<bullet> lam)" |
|
482 proof (induct lam rule: lam_induct) |
|
483 case (1 name pi) |
|
484 show "P a (pi \<bullet> Var name)" |
|
485 apply (simp) |
|
486 apply (rule a1) |
|
487 done |
|
488 next |
|
489 case (2 lam1 lam2 pi) |
|
490 have b1: "\<And>pi a. P a (pi \<bullet> lam1)" by fact |
|
491 have b2: "\<And>pi a. P a (pi \<bullet> lam2)" by fact |
|
492 show "P a (pi \<bullet> App lam1 lam2)" |
|
493 apply (simp) |
|
494 apply (rule a2) |
|
495 apply (rule b1) |
|
496 apply (rule b2) |
|
497 done |
|
498 next |
|
499 case (3 name lam pi a) |
|
500 have b: "\<And>pi a. P a (pi \<bullet> lam)" by fact |
|
501 obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)" |
|
502 apply(rule obtain_atom) |
|
503 apply(auto) |
|
504 sorry |
|
505 from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" |
|
506 apply - |
|
507 apply(rule a3) |
|
508 apply(blast) |
|
509 apply(simp add: fresh_Pair) |
|
510 done |
|
511 have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)" |
|
512 apply(rule swap_fresh_fresh) |
|
513 using fr |
|
514 apply(simp add: fresh_lam fresh_Pair) |
|
515 apply(simp add: fresh_lam fresh_Pair) |
|
516 done |
|
517 show "P a (pi \<bullet> Lam name lam)" |
|
518 apply (simp) |
|
519 apply(subst eq[symmetric]) |
|
520 using p |
|
521 apply(simp only: permute_lam) |
|
522 apply(simp add: flip_def) |
|
523 done |
|
524 qed |
|
525 then have "P a (0 \<bullet> lam)" by blast |
|
526 then show "P a lam" by simp |
|
527 qed |
|
528 |
|
529 |
|
530 lemma var_fresh: |
|
531 fixes a::"name" |
|
532 shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)" |
|
533 apply(simp add: fresh_def) |
|
534 apply(simp add: var_supp1) |
|
535 done |
|
536 |
|
537 |
|
538 |
|
539 end |
|
540 |