1 theory LetSimple2 |
|
2 imports "../Nominal2" |
|
3 begin |
|
4 |
|
5 |
|
6 atom_decl name |
|
7 |
|
8 nominal_datatype trm = |
|
9 Var "name" |
|
10 | App "trm" "trm" |
|
11 | Let as::"assn" t::"trm" binds "bn as" in t |
|
12 and assn = |
|
13 Assn "name" "trm" |
|
14 binder |
|
15 bn |
|
16 where |
|
17 "bn (Assn x t) = [atom x]" |
|
18 |
|
19 print_theorems |
|
20 |
|
21 thm bn_raw.simps |
|
22 thm permute_bn_raw.simps |
|
23 thm trm_assn.perm_bn_alpha |
|
24 thm trm_assn.permute_bn |
|
25 |
|
26 thm trm_assn.fv_defs |
|
27 thm trm_assn.eq_iff |
|
28 thm trm_assn.bn_defs |
|
29 thm trm_assn.bn_inducts |
|
30 thm trm_assn.perm_simps |
|
31 thm trm_assn.induct |
|
32 thm trm_assn.inducts |
|
33 thm trm_assn.distinct |
|
34 thm trm_assn.supp |
|
35 thm trm_assn.fresh |
|
36 thm trm_assn.exhaust |
|
37 thm trm_assn.strong_exhaust |
|
38 thm trm_assn.perm_bn_simps |
|
39 |
|
40 thm alpha_bn_raw.cases |
|
41 thm trm_assn.alpha_refl |
|
42 thm trm_assn.alpha_sym |
|
43 thm trm_assn.alpha_trans |
|
44 |
|
45 lemmas alpha_bn_cases[consumes 1] = alpha_bn_raw.cases[quot_lifted] |
|
46 |
|
47 lemma alpha_bn_refl: "alpha_bn x x" |
|
48 by(rule trm_assn.alpha_refl) |
|
49 |
|
50 lemma alpha_bn_sym: "alpha_bn x y \<Longrightarrow> alpha_bn y x" |
|
51 by (rule trm_assn.alpha_sym) |
|
52 |
|
53 lemma alpha_bn_trans: "alpha_bn x y \<Longrightarrow> alpha_bn y z \<Longrightarrow> alpha_bn x z" |
|
54 using trm_assn.alpha_trans by metis |
|
55 |
|
56 lemma fv_bn_finite[simp]: |
|
57 "finite (fv_bn as)" |
|
58 apply(case_tac as rule: trm_assn.exhaust(2)) |
|
59 apply(simp add: trm_assn.supp finite_supp) |
|
60 done |
|
61 |
|
62 |
|
63 lemma k: "A \<Longrightarrow> A \<and> A" by blast |
|
64 |
|
65 |
|
66 |
|
67 section {* definition with helper functions *} |
|
68 |
|
69 function |
|
70 apply_assn |
|
71 where |
|
72 "apply_assn f (Assn x t) = (f t)" |
|
73 apply(case_tac x) |
|
74 apply(simp) |
|
75 apply(case_tac b rule: trm_assn.exhaust(2)) |
|
76 apply(blast) |
|
77 apply(simp) |
|
78 done |
|
79 |
|
80 termination |
|
81 by lexicographic_order |
|
82 |
|
83 function |
|
84 apply_assn2 |
|
85 where |
|
86 "apply_assn2 f (Assn x t) = Assn x (f t)" |
|
87 apply(case_tac x) |
|
88 apply(simp) |
|
89 apply(case_tac b rule: trm_assn.exhaust(2)) |
|
90 apply(blast) |
|
91 apply(simp) |
|
92 done |
|
93 |
|
94 termination |
|
95 by lexicographic_order |
|
96 |
|
97 lemma [eqvt]: |
|
98 shows "p \<bullet> (apply_assn f as) = apply_assn (p \<bullet> f) (p \<bullet> as)" |
|
99 apply(induct f as rule: apply_assn.induct) |
|
100 apply(simp) |
|
101 apply(perm_simp) |
|
102 apply(rule) |
|
103 done |
|
104 |
|
105 lemma [eqvt]: |
|
106 shows "p \<bullet> (apply_assn2 f as) = apply_assn2 (p \<bullet> f) (p \<bullet> as)" |
|
107 apply(induct f as rule: apply_assn.induct) |
|
108 apply(simp) |
|
109 apply(perm_simp) |
|
110 apply(rule) |
|
111 done |
|
112 |
|
113 |
|
114 nominal_primrec |
|
115 height_trm :: "trm \<Rightarrow> nat" |
|
116 where |
|
117 "height_trm (Var x) = 1" |
|
118 | "height_trm (App l r) = max (height_trm l) (height_trm r)" |
|
119 | "height_trm (Let as b) = max (apply_assn height_trm as) (height_trm b)" |
|
120 apply (simp only: eqvt_def height_trm_graph_def) |
|
121 apply (rule, perm_simp) |
|
122 apply(rule) |
|
123 apply(rule TrueI) |
|
124 apply (case_tac x rule: trm_assn.exhaust(1)) |
|
125 apply (auto simp add: alpha_bn_refl)[3] |
|
126 apply (drule_tac x="assn" in meta_spec) |
|
127 apply (drule_tac x="trm" in meta_spec) |
|
128 apply(simp add: alpha_bn_refl) |
|
129 apply(simp_all)[5] |
|
130 apply(simp) |
|
131 apply(erule conjE)+ |
|
132 apply(erule alpha_bn_cases) |
|
133 apply(simp) |
|
134 apply (subgoal_tac "height_trm_sumC b = height_trm_sumC ba") |
|
135 apply simp |
|
136 apply(simp add: trm_assn.bn_defs) |
|
137 apply(erule_tac c="()" in Abs_lst_fcb2) |
|
138 apply(simp_all add: pure_fresh fresh_star_def)[3] |
|
139 apply(simp_all add: eqvt_at_def) |
|
140 done |
|
141 |
|
142 (* assn-function prevents automatic discharge |
|
143 termination by lexicographic_order |
|
144 *) |
|
145 |
|
146 nominal_primrec |
|
147 subst_trm :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_ [_ ::= _]" [90, 90, 90] 90) |
|
148 where |
|
149 "(Var x)[y ::= s] = (if x = y then s else (Var x))" |
|
150 | "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])" |
|
151 | "(set (bn as)) \<sharp>* (y, s) \<Longrightarrow> |
|
152 (Let as t)[y ::= s] = Let (apply_assn2 (\<lambda>t. t[y ::=s]) as) (t[y ::= s])" |
|
153 apply (simp only: eqvt_def subst_trm_graph_def) |
|
154 apply (rule, perm_simp) |
|
155 apply(rule) |
|
156 apply(rule TrueI) |
|
157 apply(case_tac x) |
|
158 apply(simp) |
|
159 apply (rule_tac y="a" and c="(b,c)" in trm_assn.strong_exhaust(1)) |
|
160 apply (auto simp add: alpha_bn_refl)[3] |
|
161 apply(simp_all)[5] |
|
162 apply(simp) |
|
163 apply(erule conjE)+ |
|
164 apply(erule alpha_bn_cases) |
|
165 apply(simp) |
|
166 apply(simp add: trm_assn.bn_defs) |
|
167 apply(erule_tac c="(ya,sa)" in Abs_lst1_fcb2) |
|
168 apply(simp add: Abs_fresh_iff fresh_star_def) |
|
169 apply(simp add: fresh_star_def) |
|
170 apply(simp_all add: eqvt_at_def perm_supp_eq fresh_star_Pair)[2] |
|
171 done |
|
172 |
|
173 |
|
174 section {* direct definitions --- problems *} |
|
175 |
|
176 lemma cheat: "P" sorry |
|
177 |
|
178 definition |
|
179 "eqvt_at_bn f as \<equiv> \<forall>p. (p \<bullet> (f as)) = f (permute_bn p as)" |
|
180 |
|
181 definition |
|
182 "alpha_bn_preserve f as \<equiv> \<forall>p. f as = f (permute_bn p as)" |
|
183 |
|
184 lemma |
|
185 fixes as::"assn" |
|
186 assumes "eqvt_at f as" |
|
187 shows "eqvt_at_bn f as" |
|
188 using assms |
|
189 unfolding eqvt_at_bn_def |
|
190 apply(rule_tac allI) |
|
191 apply(drule k) |
|
192 apply(erule conjE) |
|
193 apply(subst (asm) eqvt_at_def) |
|
194 apply(simp) |
|
195 |
|
196 oops |
|
197 |
|
198 |
|
199 |
|
200 nominal_primrec |
|
201 <<<<<<< variant A |
|
202 (invariant "\<lambda>x y. case x of Inl x1 \<Rightarrow> True | Inr x2 \<Rightarrow> alpha_bn_preserve (height_assn2::assn \<Rightarrow> nat) x2") |
|
203 >>>>>>> variant B |
|
204 ####### Ancestor |
|
205 (invariant "\<lambda>x y. case x of Inl x1 \<Rightarrow> True | Inr x2 \<Rightarrow> \<forall>p. (permute_bn p x2) = x2 \<longrightarrow> (p \<bullet> y) = y") |
|
206 ======= end |
|
207 height_trm2 :: "trm \<Rightarrow> nat" |
|
208 and height_assn2 :: "assn \<Rightarrow> nat" |
|
209 where |
|
210 "height_trm2 (Var x) = 1" |
|
211 | "height_trm2 (App l r) = max (height_trm2 l) (height_trm2 r)" |
|
212 | "set (bn as) \<sharp>* fv_bn as \<Longrightarrow> height_trm2 (Let as b) = max (height_assn2 as) (height_trm2 b)" |
|
213 | "height_assn2 (Assn x t) = (height_trm2 t)" |
|
214 thm height_trm2_height_assn2_graph.intros[no_vars] |
|
215 thm height_trm2_height_assn2_graph_def |
|
216 apply (simp only: eqvt_def height_trm2_height_assn2_graph_def) |
|
217 apply (rule, perm_simp, rule) |
|
218 -- "invariant" |
|
219 apply(simp) |
|
220 <<<<<<< variant A |
|
221 apply(simp) |
|
222 apply(simp) |
|
223 apply(simp) |
|
224 apply(simp add: alpha_bn_preserve_def) |
|
225 apply(simp add: height_assn2_def) |
|
226 apply(simp add: trm_assn.perm_bn_simps) |
|
227 apply(rule allI) |
|
228 thm height_trm2_height_assn2_graph.intros[no_vars] |
|
229 thm height_trm2_height_assn2_sumC_def |
|
230 apply(rule cheat) |
|
231 apply - |
|
232 >>>>>>> variant B |
|
233 ####### Ancestor |
|
234 apply(simp) |
|
235 apply(simp) |
|
236 apply(simp) |
|
237 apply(rule cheat) |
|
238 apply - |
|
239 ======= end |
|
240 --"completeness" |
|
241 apply (case_tac x) |
|
242 apply(simp) |
|
243 apply (rule_tac y="a" and c="a" in trm_assn.strong_exhaust(1)) |
|
244 apply (auto simp add: alpha_bn_refl)[3] |
|
245 apply (drule_tac x="assn" in meta_spec) |
|
246 apply (drule_tac x="trm" in meta_spec) |
|
247 apply(simp add: alpha_bn_refl) |
|
248 apply(rotate_tac 3) |
|
249 apply(drule meta_mp) |
|
250 apply(simp add: fresh_star_def trm_assn.fresh) |
|
251 apply(simp add: fresh_def) |
|
252 apply(subst supp_finite_atom_set) |
|
253 apply(simp) |
|
254 apply(simp) |
|
255 apply(simp) |
|
256 apply (case_tac b rule: trm_assn.exhaust(2)) |
|
257 apply (auto)[1] |
|
258 apply(simp_all)[7] |
|
259 prefer 2 |
|
260 apply(simp) |
|
261 prefer 2 |
|
262 apply(simp) |
|
263 --"let case" |
|
264 apply (simp only: meta_eq_to_obj_eq[OF height_trm2_def, symmetric, unfolded fun_eq_iff]) |
|
265 apply (simp only: meta_eq_to_obj_eq[OF height_assn2_def, symmetric, unfolded fun_eq_iff]) |
|
266 apply (subgoal_tac "eqvt_at height_assn2 as") |
|
267 apply (subgoal_tac "eqvt_at height_assn2 asa") |
|
268 apply (subgoal_tac "eqvt_at height_trm2 b") |
|
269 apply (subgoal_tac "eqvt_at height_trm2 ba") |
|
270 apply (thin_tac "eqvt_at height_trm2_height_assn2_sumC (Inr as)") |
|
271 apply (thin_tac "eqvt_at height_trm2_height_assn2_sumC (Inr asa)") |
|
272 apply (thin_tac "eqvt_at height_trm2_height_assn2_sumC (Inl b)") |
|
273 apply (thin_tac "eqvt_at height_trm2_height_assn2_sumC (Inl ba)") |
|
274 defer |
|
275 apply (simp add: eqvt_at_def height_trm2_def) |
|
276 apply (simp add: eqvt_at_def height_trm2_def) |
|
277 apply (simp add: eqvt_at_def height_assn2_def) |
|
278 apply (simp add: eqvt_at_def height_assn2_def) |
|
279 apply (subgoal_tac "height_assn2 as = height_assn2 asa") |
|
280 apply (subgoal_tac "height_trm2 b = height_trm2 ba") |
|
281 apply simp |
|
282 apply(simp) |
|
283 apply(erule conjE)+ |
|
284 apply(erule alpha_bn_cases) |
|
285 apply(simp) |
|
286 apply(simp add: trm_assn.bn_defs) |
|
287 apply(erule_tac c="()" in Abs_lst_fcb2) |
|
288 apply(simp_all add: fresh_star_def pure_fresh)[3] |
|
289 apply(simp add: eqvt_at_def) |
|
290 apply(simp add: eqvt_at_def) |
|
291 apply(drule Inl_inject) |
|
292 apply(simp (no_asm_use)) |
|
293 apply(clarify) |
|
294 apply(erule alpha_bn_cases) |
|
295 apply(simp del: trm_assn.eq_iff) |
|
296 apply(simp only: trm_assn.bn_defs) |
|
297 <<<<<<< variant A |
|
298 apply(erule_tac c="()" in Abs_lst1_fcb2') |
|
299 apply(simp_all add: fresh_star_def pure_fresh)[3] |
|
300 apply(simp add: eqvt_at_bn_def) |
|
301 apply(simp add: trm_assn.perm_bn_simps) |
|
302 apply(simp add: eqvt_at_bn_def) |
|
303 apply(simp add: trm_assn.perm_bn_simps) |
|
304 done |
|
305 |
|
306 >>>>>>> variant B |
|
307 apply(erule_tac c="(trm_rawa)" in Abs_lst1_fcb2') |
|
308 apply(simp_all add: fresh_star_def pure_fresh)[2] |
|
309 apply(simp add: trm_assn.supp) |
|
310 apply(simp add: fresh_def) |
|
311 apply(subst (asm) supp_finite_atom_set) |
|
312 apply(simp add: finite_supp) |
|
313 apply(subst (asm) supp_finite_atom_set) |
|
314 apply(simp add: finite_supp) |
|
315 apply(simp) |
|
316 apply(simp add: eqvt_at_def perm_supp_eq) |
|
317 apply(simp add: eqvt_at_def perm_supp_eq) |
|
318 done |
|
319 ####### Ancestor |
|
320 apply(erule_tac c="()" in Abs_lst1_fcb2') |
|
321 apply(simp_all add: fresh_star_def pure_fresh)[3] |
|
322 |
|
323 oops |
|
324 ======= end |
|
325 |
|
326 termination by lexicographic_order |
|
327 |
|
328 lemma ww1: |
|
329 shows "finite (fv_trm t)" |
|
330 and "finite (fv_bn as)" |
|
331 apply(induct t and as rule: trm_assn.inducts) |
|
332 apply(simp_all add: trm_assn.fv_defs supp_at_base) |
|
333 done |
|
334 |
|
335 text {* works, but only because no recursion in as *} |
|
336 |
|
337 nominal_primrec (invariant "\<lambda>x (y::atom set). finite y") |
|
338 frees_set :: "trm \<Rightarrow> atom set" |
|
339 where |
|
340 "frees_set (Var x) = {atom x}" |
|
341 | "frees_set (App t1 t2) = frees_set t1 \<union> frees_set t2" |
|
342 | "frees_set (Let as t) = (frees_set t) - (set (bn as)) \<union> (fv_bn as)" |
|
343 apply(simp add: eqvt_def frees_set_graph_def) |
|
344 apply(rule, perm_simp, rule) |
|
345 apply(erule frees_set_graph.induct) |
|
346 apply(auto simp add: ww1)[3] |
|
347 apply(rule_tac y="x" in trm_assn.exhaust(1)) |
|
348 apply(auto simp add: alpha_bn_refl)[3] |
|
349 apply(drule_tac x="assn" in meta_spec) |
|
350 apply(drule_tac x="trm" in meta_spec) |
|
351 apply(simp add: alpha_bn_refl) |
|
352 apply(simp_all)[5] |
|
353 apply(simp) |
|
354 apply(erule conjE) |
|
355 apply(erule alpha_bn_cases) |
|
356 apply(simp add: trm_assn.bn_defs) |
|
357 apply(simp add: trm_assn.fv_defs) |
|
358 (* apply(erule_tac c="(trm_rawa)" in Abs_lst1_fcb2) *) |
|
359 apply(subgoal_tac " frees_set_sumC t - {atom name} = frees_set_sumC ta - {atom namea}") |
|
360 apply(simp) |
|
361 apply(erule_tac c="()" in Abs_lst1_fcb2) |
|
362 apply(simp add: fresh_minus_atom_set) |
|
363 apply(simp add: fresh_star_def fresh_Unit) |
|
364 apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) |
|
365 apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl) |
|
366 done |
|
367 |
|
368 termination |
|
369 by lexicographic_order |
|
370 |
|
371 lemma test: |
|
372 assumes a: "\<exists>y. f x = Inl y" |
|
373 shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl ((p \<bullet> f) (p \<bullet> x))" |
|
374 using a |
|
375 apply clarify |
|
376 apply(frule_tac p="p" in permute_boolI) |
|
377 apply(simp (no_asm_use) only: eqvts) |
|
378 apply(subst (asm) permute_fun_app_eq) |
|
379 back |
|
380 apply(simp) |
|
381 done |
|
382 |
|
383 |
|
384 nominal_primrec (default "sum_case (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined)") |
|
385 subst_trm2 :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_ [_ ::trm2= _]" [90, 90, 90] 90) and |
|
386 subst_assn2 :: "assn \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> assn" ("_ [_ ::assn2= _]" [90, 90, 90] 90) |
|
387 where |
|
388 "(Var x)[y ::trm2= s] = (if x = y then s else (Var x))" |
|
389 | "(App t1 t2)[y ::trm2= s] = App (t1[y ::trm2= s]) (t2[y ::trm2= s])" |
|
390 | "(set (bn as)) \<sharp>* (y, s, fv_bn as) \<Longrightarrow> (Let as t)[y ::trm2= s] = Let (ast[y ::assn2= s]) (t[y ::trm2= s])" |
|
391 | "(Assn x t)[y ::assn2= s] = Assn x (t[y ::trm2= s])" |
|
392 apply(subgoal_tac "\<And>p x r. subst_trm2_subst_assn2_graph x r \<Longrightarrow> subst_trm2_subst_assn2_graph (p \<bullet> x) (p \<bullet> r)") |
|
393 apply(simp add: eqvt_def) |
|
394 apply(rule allI) |
|
395 apply(simp add: permute_fun_def permute_bool_def) |
|
396 apply(rule ext) |
|
397 apply(rule ext) |
|
398 apply(rule iffI) |
|
399 apply(drule_tac x="p" in meta_spec) |
|
400 apply(drule_tac x="- p \<bullet> x" in meta_spec) |
|
401 apply(drule_tac x="- p \<bullet> xa" in meta_spec) |
|
402 apply(simp) |
|
403 apply(drule_tac x="-p" in meta_spec) |
|
404 apply(drule_tac x="x" in meta_spec) |
|
405 apply(drule_tac x="xa" in meta_spec) |
|
406 apply(simp) |
|
407 --"Eqvt One way" |
|
408 defer |
|
409 apply(rule TrueI) |
|
410 apply(case_tac x) |
|
411 apply(simp) |
|
412 apply(case_tac a) |
|
413 apply(simp) |
|
414 apply(rule_tac y="aa" and c="(b, c, aa)" in trm_assn.strong_exhaust(1)) |
|
415 apply(blast)+ |
|
416 apply(simp) |
|
417 apply(drule_tac x="assn" in meta_spec) |
|
418 apply(drule_tac x="b" in meta_spec) |
|
419 apply(drule_tac x="c" in meta_spec) |
|
420 apply(drule_tac x="trm" in meta_spec) |
|
421 apply(simp add: trm_assn.alpha_refl) |
|
422 apply(rotate_tac 5) |
|
423 apply(drule meta_mp) |
|
424 apply(simp add: fresh_star_Pair) |
|
425 apply(simp add: fresh_star_def trm_assn.fresh) |
|
426 apply(simp add: fresh_def) |
|
427 apply(subst supp_finite_atom_set) |
|
428 apply(simp) |
|
429 apply(simp) |
|
430 apply(simp) |
|
431 apply(case_tac b) |
|
432 apply(simp) |
|
433 apply(rule_tac y="a" in trm_assn.exhaust(2)) |
|
434 apply(simp) |
|
435 apply(blast) |
|
436 --"compatibility" |
|
437 apply(all_trivials) |
|
438 apply(simp) |
|
439 apply(simp) |
|
440 prefer 2 |
|
441 apply(simp) |
|
442 apply(drule Inl_inject) |
|
443 apply(rule arg_cong) |
|
444 back |
|
445 apply (simp only: meta_eq_to_obj_eq[OF subst_trm2_def, symmetric, unfolded fun_eq_iff]) |
|
446 apply (simp only: meta_eq_to_obj_eq[OF subst_assn2_def, symmetric, unfolded fun_eq_iff]) |
|
447 apply (subgoal_tac "eqvt_at (\<lambda>ast. subst_assn2 ast ya sa) ast") |
|
448 apply (subgoal_tac "eqvt_at (\<lambda>asta. subst_assn2 asta ya sa) asta") |
|
449 apply (subgoal_tac "eqvt_at (\<lambda>t. subst_trm2 t ya sa) t") |
|
450 apply (subgoal_tac "eqvt_at (\<lambda>ta. subst_trm2 ta ya sa) ta") |
|
451 apply (thin_tac "eqvt_at subst_trm2_subst_assn2_sumC (Inr (ast, y, s))") |
|
452 apply (thin_tac "eqvt_at subst_trm2_subst_assn2_sumC (Inr (asta, ya, sa))") |
|
453 apply (thin_tac "eqvt_at subst_trm2_subst_assn2_sumC (Inl (t, y, s))") |
|
454 apply (thin_tac "eqvt_at subst_trm2_subst_assn2_sumC (Inl (ta, ya, sa))") |
|
455 apply(simp) |
|
456 (* HERE *) |
|
457 apply (subgoal_tac "subst_assn2 ast y s= subst_assn2 asta ya sa") |
|
458 apply (subgoal_tac "subst_trm2 t y s = subst_trm2 ta ya sa") |
|
459 apply(simp) |
|
460 apply(simp) |
|
461 apply(erule_tac conjE)+ |
|
462 apply(erule alpha_bn_cases) |
|
463 apply(simp add: trm_assn.bn_defs) |
|
464 apply(rotate_tac 7) |
|
465 apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) |
|
466 apply(erule fresh_eqvt_at) |
|
467 |
|
468 |
|
469 thm fresh_eqvt_at |
|
470 apply(simp add: Abs_fresh_iff) |
|
471 apply(simp add: fresh_star_def fresh_Pair) |
|
472 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
473 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
474 |
|
475 |
|
476 |
|
477 apply(simp_all add: fresh_star_def fresh_Pair_elim)[1] |
|
478 apply(blast) |
|
479 apply(simp_all)[5] |
|
480 apply(simp (no_asm_use)) |
|
481 apply(simp) |
|
482 apply(erule conjE)+ |
|
483 apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) |
|
484 apply(simp add: Abs_fresh_iff) |
|
485 apply(simp add: fresh_star_def fresh_Pair) |
|
486 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
487 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
488 done |
|
489 |
|
490 |
|
491 end |
|