|
1 |
|
2 |
|
3 theory Fresh = Main + Swap + Terms + Disagreement: |
|
4 |
|
5 types fresh_envs = "(string \<times> string)set" |
|
6 |
|
7 consts |
|
8 fresh :: "(fresh_envs \<times> string \<times> trm) set" |
|
9 syntax |
|
10 "_fresh_judge" :: "fresh_envs \<Rightarrow> string \<Rightarrow> trm \<Rightarrow> bool" (" _ \<turnstile> _ \<sharp> _" [80,80,80] 80) |
|
11 translations |
|
12 "nabla \<turnstile> a \<sharp> t" \<rightleftharpoons> "(nabla,a,t) \<in> fresh" |
|
13 |
|
14 inductive fresh |
|
15 intros |
|
16 fresh_abst_ab[intro!]: "\<lbrakk>(nabla\<turnstile>a\<sharp>t); a\<noteq>b\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>Abst b t)" |
|
17 fresh_abst_aa[intro!]: "(nabla\<turnstile>a\<sharp>Abst a t)" |
|
18 fresh_unit[intro!]: "(nabla\<turnstile>a\<sharp>Unit)" |
|
19 fresh_atom[intro!]: "a\<noteq>b \<Longrightarrow> (nabla\<turnstile>a\<sharp>Atom b)" |
|
20 fresh_susp[intro!]: "(swapas (rev pi) a,X)\<in>nabla \<Longrightarrow> (nabla\<turnstile>a\<sharp>Susp pi X)" |
|
21 fresh_paar[intro!]: "\<lbrakk>(nabla\<turnstile>a\<sharp>t1);(nabla\<turnstile>a\<sharp>t2)\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>Paar t1 t2)" |
|
22 fresh_func[intro!]: "(nabla\<turnstile>a\<sharp>t) \<Longrightarrow> (nabla\<turnstile>a\<sharp>Func F t)" |
|
23 |
|
24 lemma fresh_atom_elim[elim!]: "(nabla\<turnstile>a\<sharp>Atom b) \<Longrightarrow> a\<noteq>b" |
|
25 apply(ind_cases "(nabla \<turnstile>a\<sharp>Atom b)") |
|
26 apply(auto) |
|
27 done |
|
28 lemma fresh_susp_elim[elim!]: "(nabla\<turnstile>a\<sharp>Susp pi X) \<Longrightarrow> (swapas (rev pi) a,X)\<in>nabla" |
|
29 apply(ind_cases "(nabla\<turnstile>a\<sharp>Susp pi X)") |
|
30 apply(auto) |
|
31 done |
|
32 lemma fresh_paar_elim[elim!]: "(nabla\<turnstile>a\<sharp>Paar t1 t2) \<Longrightarrow> (nabla\<turnstile>a\<sharp>t1)\<and>(nabla \<turnstile>a\<sharp>t2)" |
|
33 apply(ind_cases "(nabla\<turnstile>a\<sharp>Paar t1 t2)") |
|
34 apply(auto) |
|
35 done |
|
36 lemma fresh_func_elim[elim!]: "(nabla\<turnstile>a\<sharp>Func F t) \<Longrightarrow> (nabla\<turnstile>a\<sharp>t)" |
|
37 apply(ind_cases "nabla\<turnstile>a\<sharp>Func F t") |
|
38 apply(auto) |
|
39 done |
|
40 lemma fresh_abst_ab_elim[elim!]: "\<lbrakk>nabla\<turnstile>a\<sharp>Abst b t;a\<noteq>b\<rbrakk> \<Longrightarrow> (nabla\<turnstile>a\<sharp>t)" |
|
41 apply(ind_cases "nabla\<turnstile>a\<sharp>Abst b t", auto) |
|
42 done |
|
43 |
|
44 lemma fresh_swap_left: "(nabla\<turnstile>a\<sharp>swap pi t) \<longrightarrow> (nabla\<turnstile>swapas (rev pi) a\<sharp>t)" |
|
45 apply(induct t) |
|
46 apply(simp_all) |
|
47 -- Abst |
|
48 apply(rule impI) |
|
49 apply(case_tac "swapas (rev pi) a = list") |
|
50 apply(force) |
|
51 apply(force dest!: fresh_abst_ab_elim) |
|
52 --Susp |
|
53 apply(force dest!: fresh_susp_elim simp add: swapas_append[THEN sym]) |
|
54 --Unit |
|
55 apply(force) |
|
56 --Atom |
|
57 apply(force dest!: fresh_atom_elim) |
|
58 --Paar |
|
59 apply(force dest!: fresh_paar_elim) |
|
60 -- Func |
|
61 apply(force dest!: fresh_func_elim) |
|
62 done |
|
63 |
|
64 lemma fresh_swap_right: "(nabla\<turnstile>swapas (rev pi) a\<sharp>t) \<longrightarrow> (nabla\<turnstile>a\<sharp>swap pi t)" |
|
65 apply(induct t) |
|
66 apply(simp_all) |
|
67 -- Abst |
|
68 apply(rule impI) |
|
69 apply(case_tac "a = swapas pi list") |
|
70 apply(force) |
|
71 apply(force dest!: fresh_abst_ab_elim) |
|
72 --Susp |
|
73 apply(force dest!: fresh_susp_elim simp add: swapas_append[THEN sym]) |
|
74 --Unit |
|
75 apply(force) |
|
76 --Atom |
|
77 apply(force dest!: fresh_atom_elim) |
|
78 --Paar |
|
79 apply(force dest!: fresh_paar_elim) |
|
80 -- Func |
|
81 apply(force dest!: fresh_func_elim) |
|
82 done |
|
83 |
|
84 lemma fresh_weak: "nabla1\<turnstile>a\<sharp>t\<longrightarrow>(nabla1\<union>nabla2)\<turnstile>a\<sharp>t" |
|
85 apply(rule impI) |
|
86 apply(erule fresh.induct) |
|
87 apply(auto) |
|
88 done |
|
89 |
|
90 end |