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