1 theory Term8 |
|
2 imports "../../Nominal-General/Nominal2_Atoms" "../../Nominal-General/Nominal2_Eqvt" "../../Nominal-General/Nominal2_Supp" "../Abs" "../Perm" "../Fv" "../Rsp" "../Lift" "../../Attic/Prove" |
|
3 begin |
|
4 |
|
5 atom_decl name |
|
6 |
|
7 datatype rfoo8 = |
|
8 Foo0 "name" |
|
9 | Foo1 "rbar8" "rfoo8" --"bind bv(bar) in foo" |
|
10 and rbar8 = |
|
11 Bar0 "name" |
|
12 | Bar1 "name" "name" "rbar8" --"bind second name in b" |
|
13 |
|
14 primrec |
|
15 rbv8 |
|
16 where |
|
17 "rbv8 (Bar0 x) = {}" |
|
18 | "rbv8 (Bar1 v x b) = {atom v}" |
|
19 |
|
20 setup {* snd o define_raw_perms (Datatype.the_info @{theory} "Term8.rfoo8") 2 *} |
|
21 print_theorems |
|
22 |
|
23 ML define_fv_alpha_export |
|
24 local_setup {* snd o define_fv_alpha_export (Datatype.the_info @{theory} "Term8.rfoo8") [ |
|
25 [[], [(SOME (@{term rbv8}, false), 0, 1, AlphaGen)]], [[], [(NONE, 1, 2, AlphaGen)]]] |
|
26 [(@{term "rbv8"}, 1, [[], [(0, NONE), (2, SOME @{term rbv8})]])] *} |
|
27 notation |
|
28 alpha_rfoo8 ("_ \<approx>f' _" [100, 100] 100) and |
|
29 alpha_rbar8 ("_ \<approx>b' _" [100, 100] 100) |
|
30 thm alpha_rfoo8_alpha_rbar8_alpha_rbv8.intros[no_vars] |
|
31 thm fv_rbar8.simps fv_rfoo8_fv_rbv8.simps |
|
32 |
|
33 inductive |
|
34 alpha8f :: "rfoo8 \<Rightarrow> rfoo8 \<Rightarrow> bool" ("_ \<approx>f _" [100, 100] 100) |
|
35 and |
|
36 alpha8b :: "rbar8 \<Rightarrow> rbar8 \<Rightarrow> bool" ("_ \<approx>b _" [100, 100] 100) |
|
37 and |
|
38 alpha8bv:: "rbar8 \<Rightarrow> rbar8 \<Rightarrow> bool" ("_ \<approx>bv _" [100, 100] 100) |
|
39 where |
|
40 "name = namea \<Longrightarrow> Foo0 name \<approx>f Foo0 namea" |
|
41 | "\<exists>pi. rbar8 \<approx>bv rbar8a \<and> |
|
42 (rbv8 rbar8, rfoo8) \<approx>gen alpha8f fv_rfoo8 pi (rbv8 rbar8a, rfoo8a) \<Longrightarrow> |
|
43 Foo1 rbar8 rfoo8 \<approx>f Foo1 rbar8a rfoo8a" |
|
44 | "name = namea \<Longrightarrow> Bar0 name \<approx>b Bar0 namea" |
|
45 | "\<exists>pi. name1 = name1a \<and> ({atom name2}, rbar8) \<approx>gen alpha8b fv_rbv8 pi ({atom name2a}, rbar8a) \<Longrightarrow> |
|
46 Bar1 name1 name2 rbar8 \<approx>b Bar1 name1a name2a rbar8a" |
|
47 | "name = namea \<Longrightarrow> alpha8bv (Bar0 name) (Bar0 namea)" |
|
48 | "({atom name2}, rbar8) \<approx>gen alpha8b fv_rbv8 pi ({atom name2a}, rbar8a) \<Longrightarrow> |
|
49 alpha8bv (Bar1 name1 name2 rbar8) (Bar1 name1a name2a rbar8a)" |
|
50 |
|
51 lemma "(alpha8b ===> op =) rbv8 rbv8" |
|
52 apply rule |
|
53 apply (induct_tac a b rule: alpha8f_alpha8b_alpha8bv.inducts(2)) |
|
54 apply (simp_all) |
|
55 done |
|
56 |
|
57 lemma fv_rbar8_rsp_hlp: "x \<approx>b y \<Longrightarrow> fv_rbar8 x = fv_rbar8 y" |
|
58 apply (erule alpha8f_alpha8b_alpha8bv.inducts(2)) |
|
59 apply (simp_all add: alpha_gen) |
|
60 apply clarify |
|
61 sorry |
|
62 |
|
63 lemma "(alpha8b ===> op =) fv_rbar8 fv_rbar8" |
|
64 apply simp apply clarify apply (simp add: fv_rbar8_rsp_hlp) |
|
65 done |
|
66 |
|
67 lemma "(alpha8f ===> op =) fv_rfoo8 fv_rfoo8" |
|
68 apply simp apply clarify |
|
69 apply (erule alpha8f_alpha8b_alpha8bv.inducts(1)) |
|
70 apply (simp_all add: alpha_gen fv_rbar8_rsp_hlp) |
|
71 |
|
72 done |
|
73 |
|
74 end |
|