| author | haftmann |
| Mon, 21 May 2007 19:11:42 +0200 | |
| changeset 23063 | b4ee6ec4f9c6 |
| parent 22730 | 8bcc8809ed3b |
| child 23142 | cb1dbe64a4d5 |
| permissions | -rw-r--r-- |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
1 |
(* $Id$ *) |
| 18106 | 2 |
|
| 19496 | 3 |
theory SN |
|
21107
e69c0e82955a
new file for defining functions in the lambda-calculus
urbanc
parents:
19972
diff
changeset
|
4 |
imports Lam_Funs |
| 18106 | 5 |
begin |
6 |
||
| 18269 | 7 |
text {* Strong Normalisation proof from the Proofs and Types book *}
|
| 18106 | 8 |
|
9 |
section {* Beta Reduction *}
|
|
10 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
11 |
lemma subst_rename: |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
12 |
assumes a: "c\<sharp>t1" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
13 |
shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
14 |
using a |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
15 |
by (nominal_induct t1 avoiding: a c t2 rule: lam.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
16 |
(auto simp add: calc_atm fresh_atm abs_fresh) |
| 18106 | 17 |
|
| 18313 | 18 |
lemma forget: |
19 |
assumes a: "a\<sharp>t1" |
|
20 |
shows "t1[a::=t2] = t1" |
|
21 |
using a |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
22 |
by (nominal_induct t1 avoiding: a t2 rule: lam.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
23 |
(auto simp add: abs_fresh fresh_atm) |
| 18106 | 24 |
|
| 18313 | 25 |
lemma fresh_fact: |
26 |
fixes a::"name" |
|
27 |
assumes a: "a\<sharp>t1" |
|
28 |
and b: "a\<sharp>t2" |
|
| 22540 | 29 |
shows "a\<sharp>t1[b::=t2]" |
| 18313 | 30 |
using a b |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
31 |
by (nominal_induct t1 avoiding: a b t2 rule: lam.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
32 |
(auto simp add: abs_fresh fresh_atm) |
| 18106 | 33 |
|
| 22540 | 34 |
lemma fresh_fact': |
35 |
fixes a::"name" |
|
36 |
assumes a: "a\<sharp>t2" |
|
37 |
shows "a\<sharp>t1[a::=t2]" |
|
38 |
using a |
|
39 |
by (nominal_induct t1 avoiding: a t2 rule: lam.induct) |
|
40 |
(auto simp add: abs_fresh fresh_atm) |
|
41 |
||
| 18383 | 42 |
lemma subst_lemma: |
| 18313 | 43 |
assumes a: "x\<noteq>y" |
44 |
and b: "x\<sharp>L" |
|
45 |
shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
46 |
using a b |
|
|
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18654
diff
changeset
|
47 |
by (nominal_induct M avoiding: x y N L rule: lam.induct) |
| 18313 | 48 |
(auto simp add: fresh_fact forget) |
| 18106 | 49 |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
50 |
lemma id_subs: |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
51 |
shows "t[x::=Var x] = t" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
52 |
by (nominal_induct t avoiding: x rule: lam.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
53 |
(simp_all add: fresh_atm) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
54 |
|
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
55 |
lemma subst_eqvt[eqvt]: |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
56 |
fixes pi::"name prm" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
57 |
and t::"lam" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
58 |
shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]" |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
59 |
by (nominal_induct t avoiding: x t' rule: lam.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
60 |
(perm_simp add: fresh_bij)+ |
| 18106 | 61 |
|
| 22271 | 62 |
inductive2 Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
|
63 |
where |
|
| 18106 | 64 |
b1[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App s1 t)\<longrightarrow>\<^isub>\<beta>(App s2 t)" |
| 22271 | 65 |
| b2[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (App t s1)\<longrightarrow>\<^isub>\<beta>(App t s2)" |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
66 |
| b3[intro!]: "s1\<longrightarrow>\<^isub>\<beta>s2 \<Longrightarrow> (Lam [a].s1)\<longrightarrow>\<^isub>\<beta> (Lam [a].s2)" |
| 22540 | 67 |
| b4[intro!]: "a\<sharp>s2 \<Longrightarrow>(App (Lam [a].s1) s2)\<longrightarrow>\<^isub>\<beta>(s1[a::=s2])" |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
68 |
|
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
69 |
equivariance Beta |
|
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
70 |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
71 |
nominal_inductive Beta |
| 22540 | 72 |
by (simp_all add: abs_fresh fresh_fact') |
| 18106 | 73 |
|
| 22271 | 74 |
abbreviation "Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80) where
|
75 |
"t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2 \<equiv> Beta\<^sup>*\<^sup>* t1 t2" |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
76 |
|
| 18378 | 77 |
lemma supp_beta: |
78 |
assumes a: "t\<longrightarrow>\<^isub>\<beta> s" |
|
79 |
shows "(supp s)\<subseteq>((supp t)::name set)" |
|
80 |
using a |
|
81 |
by (induct) |
|
82 |
(auto intro!: simp add: abs_supp lam.supp subst_supp) |
|
83 |
||
| 18106 | 84 |
lemma beta_abs: "Lam [a].t\<longrightarrow>\<^isub>\<beta> t'\<Longrightarrow>\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^isub>\<beta> t''" |
| 22271 | 85 |
apply(ind_cases2 "Lam [a].t \<longrightarrow>\<^isub>\<beta> t'") |
| 18106 | 86 |
apply(auto simp add: lam.distinct lam.inject) |
87 |
apply(auto simp add: alpha) |
|
88 |
apply(rule_tac x="[(a,aa)]\<bullet>s2" in exI) |
|
89 |
apply(rule conjI) |
|
90 |
apply(rule sym) |
|
91 |
apply(rule pt_bij2[OF pt_name_inst, OF at_name_inst]) |
|
92 |
apply(simp) |
|
93 |
apply(rule pt_name3) |
|
94 |
apply(simp add: at_ds5[OF at_name_inst]) |
|
95 |
apply(rule conjI) |
|
96 |
apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] calc_atm) |
|
97 |
apply(force dest!: supp_beta simp add: fresh_def) |
|
|
22541
c33b542394f3
the name for the collection of equivariance lemmas is now eqvts (changed from eqvt) in order to avoid clashes with eqvt-lemmas generated in nominal_inductive
urbanc
parents:
22540
diff
changeset
|
98 |
apply(force intro!: eqvts) |
| 18106 | 99 |
done |
100 |
||
| 18313 | 101 |
lemma beta_subst: |
| 18106 | 102 |
assumes a: "M \<longrightarrow>\<^isub>\<beta> M'" |
103 |
shows "M[x::=N]\<longrightarrow>\<^isub>\<beta> M'[x::=N]" |
|
104 |
using a |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
105 |
apply(nominal_induct M M' avoiding: x N rule: Beta.strong_induct) |
|
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
106 |
apply(auto simp add: fresh_atm subst_lemma fresh_fact) |
| 18106 | 107 |
done |
108 |
||
| 18383 | 109 |
section {* types *}
|
110 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
111 |
nominal_datatype ty = |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
112 |
TVar "nat" |
| 18106 | 113 |
| TArr "ty" "ty" (infix "\<rightarrow>" 200) |
114 |
||
115 |
lemma perm_ty[simp]: |
|
116 |
fixes pi ::"name prm" |
|
117 |
and \<tau> ::"ty" |
|
118 |
shows "pi\<bullet>\<tau> = \<tau>" |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
119 |
by (nominal_induct \<tau> rule: ty.induct) |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
120 |
(simp_all add: perm_nat_def) |
| 18106 | 121 |
|
122 |
lemma fresh_ty: |
|
123 |
fixes a ::"name" |
|
124 |
and \<tau> ::"ty" |
|
125 |
shows "a\<sharp>\<tau>" |
|
126 |
by (simp add: fresh_def supp_def) |
|
127 |
||
128 |
(* valid contexts *) |
|
129 |
||
130 |
consts |
|
131 |
"dom_ty" :: "(name\<times>ty) list \<Rightarrow> (name list)" |
|
132 |
primrec |
|
133 |
"dom_ty [] = []" |
|
134 |
"dom_ty (x#\<Gamma>) = (fst x)#(dom_ty \<Gamma>)" |
|
135 |
||
| 22271 | 136 |
inductive2 valid :: "(name\<times>ty) list \<Rightarrow> bool" |
137 |
where |
|
138 |
v1[intro]: "valid []" |
|
139 |
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)" |
|
| 18106 | 140 |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
141 |
equivariance valid |
| 18106 | 142 |
|
143 |
(* typing judgements *) |
|
144 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
145 |
lemma fresh_context: |
| 18106 | 146 |
fixes \<Gamma> :: "(name\<times>ty)list" |
147 |
and a :: "name" |
|
| 18378 | 148 |
assumes a: "a\<sharp>\<Gamma>" |
149 |
shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)" |
|
150 |
using a |
|
151 |
apply(induct \<Gamma>) |
|
| 18106 | 152 |
apply(auto simp add: fresh_prod fresh_list_cons fresh_atm) |
153 |
done |
|
154 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
155 |
inductive_cases2 valid_elim[elim]: "valid ((a,\<tau>)#\<Gamma>)" |
| 18106 | 156 |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
157 |
lemma valid_unicity: |
| 18378 | 158 |
assumes a: "valid \<Gamma>" |
159 |
and b: "(c,\<sigma>)\<in>set \<Gamma>" |
|
160 |
and c: "(c,\<tau>)\<in>set \<Gamma>" |
|
161 |
shows "\<sigma>=\<tau>" |
|
162 |
using a b c |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
163 |
by (induct \<Gamma>) (auto dest!: fresh_context) |
| 18106 | 164 |
|
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
165 |
inductive2 typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ : _" [60,60,60] 60)
|
| 22271 | 166 |
where |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
167 |
t1[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>" |
|
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
168 |
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>" |
| 22271 | 169 |
| t3[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>" |
| 18106 | 170 |
|
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
171 |
equivariance typing |
|
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
172 |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
173 |
nominal_inductive typing |
|
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
174 |
by (simp_all add: abs_fresh fresh_ty) |
| 18106 | 175 |
|
|
21107
e69c0e82955a
new file for defining functions in the lambda-calculus
urbanc
parents:
19972
diff
changeset
|
176 |
abbreviation |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
177 |
"sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ \<subseteq> _" [60,60] 60)
|
|
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
178 |
where |
|
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
179 |
"\<Gamma>1 \<subseteq> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>2" |
| 18106 | 180 |
|
| 18313 | 181 |
lemma weakening: |
182 |
assumes a: "\<Gamma>1 \<turnstile> t : \<sigma>" |
|
183 |
and b: "valid \<Gamma>2" |
|
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
184 |
and c: "\<Gamma>1 \<subseteq> \<Gamma>2" |
|
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
185 |
shows "\<Gamma>2 \<turnstile> t : \<sigma>" |
| 18313 | 186 |
using a b c |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
187 |
apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing.strong_induct) |
|
21107
e69c0e82955a
new file for defining functions in the lambda-calculus
urbanc
parents:
19972
diff
changeset
|
188 |
apply(auto | atomize)+ |
| 18106 | 189 |
done |
190 |
||
| 18378 | 191 |
lemma in_ctxt: |
192 |
assumes a: "(a,\<tau>)\<in>set \<Gamma>" |
|
193 |
shows "a\<in>set(dom_ty \<Gamma>)" |
|
194 |
using a |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
195 |
by (induct \<Gamma>) (auto) |
| 18106 | 196 |
|
197 |
lemma free_vars: |
|
198 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
|
199 |
shows " (supp t)\<subseteq>set(dom_ty \<Gamma>)" |
|
200 |
using a |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
201 |
by (nominal_induct \<Gamma> t \<tau> rule: typing.strong_induct) |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
202 |
(auto simp add: lam.supp abs_supp supp_atm in_ctxt) |
| 18106 | 203 |
|
204 |
lemma t1_elim: "\<Gamma> \<turnstile> Var a : \<tau> \<Longrightarrow> valid \<Gamma> \<and> (a,\<tau>) \<in> set \<Gamma>" |
|
| 22271 | 205 |
apply(ind_cases2 "\<Gamma> \<turnstile> Var a : \<tau>") |
| 18106 | 206 |
apply(auto simp add: lam.inject lam.distinct) |
207 |
done |
|
208 |
||
209 |
lemma t2_elim: "\<Gamma> \<turnstile> App t1 t2 : \<sigma> \<Longrightarrow> \<exists>\<tau>. (\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma> \<and> \<Gamma> \<turnstile> t2 : \<tau>)" |
|
| 22271 | 210 |
apply(ind_cases2 "\<Gamma> \<turnstile> App t1 t2 : \<sigma>") |
| 18106 | 211 |
apply(auto simp add: lam.inject lam.distinct) |
212 |
done |
|
213 |
||
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
214 |
lemma t3_elim: "\<lbrakk>\<Gamma> \<turnstile> Lam [a].t : \<sigma>;a\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> \<exists>\<tau> \<tau>'. \<sigma>=\<tau>\<rightarrow>\<tau>' \<and> (a,\<tau>)#\<Gamma> \<turnstile> t : \<tau>'" |
| 22271 | 215 |
apply(ind_cases2 "\<Gamma> \<turnstile> Lam [a].t : \<sigma>") |
| 18106 | 216 |
apply(auto simp add: lam.distinct lam.inject alpha) |
| 22542 | 217 |
apply(drule_tac pi="[(a,aa)]::name prm" in typing.eqvt) |
| 18106 | 218 |
apply(simp) |
219 |
apply(subgoal_tac "([(a,aa)]::name prm)\<bullet>\<Gamma> = \<Gamma>")(*A*) |
|
220 |
apply(force simp add: calc_atm) |
|
221 |
(*A*) |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
222 |
apply(force intro!: perm_fresh_fresh) |
| 18106 | 223 |
done |
224 |
||
225 |
lemma typing_valid: |
|
226 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
|
227 |
shows "valid \<Gamma>" |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
228 |
using a by (induct, auto) |
| 18106 | 229 |
|
| 18378 | 230 |
lemma ty_subs: |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
231 |
assumes a: "(c,\<sigma>)#\<Gamma> \<turnstile> t1:\<tau>" |
| 18378 | 232 |
and b: "\<Gamma>\<turnstile> t2:\<sigma>" |
233 |
shows "\<Gamma> \<turnstile> t1[c::=t2]:\<tau>" |
|
234 |
using a b |
|
|
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18654
diff
changeset
|
235 |
proof(nominal_induct t1 avoiding: \<Gamma> \<sigma> \<tau> c t2 rule: lam.induct) |
| 18313 | 236 |
case (Var a) |
| 18378 | 237 |
have a1: "\<Gamma> \<turnstile>t2:\<sigma>" by fact |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
238 |
have a2: "(c,\<sigma>)#\<Gamma> \<turnstile> Var a : \<tau>" by fact |
| 18378 | 239 |
hence a21: "(a,\<tau>)\<in>set((c,\<sigma>)#\<Gamma>)" and a22: "valid((c,\<sigma>)#\<Gamma>)" by (auto dest: t1_elim) |
240 |
from a22 have a23: "valid \<Gamma>" and a24: "c\<sharp>\<Gamma>" by (auto dest: valid_elim) |
|
241 |
from a24 have a25: "\<not>(\<exists>\<tau>. (c,\<tau>)\<in>set \<Gamma>)" by (rule fresh_context) |
|
242 |
show "\<Gamma>\<turnstile>(Var a)[c::=t2] : \<tau>" |
|
243 |
proof (cases "a=c", simp_all) |
|
244 |
assume case1: "a=c" |
|
245 |
show "\<Gamma> \<turnstile> t2:\<tau>" using a1 |
|
246 |
proof (cases "\<sigma>=\<tau>") |
|
247 |
assume "\<sigma>=\<tau>" thus ?thesis using a1 by simp |
|
248 |
next |
|
249 |
assume a3: "\<sigma>\<noteq>\<tau>" |
|
250 |
show ?thesis |
|
251 |
proof (rule ccontr) |
|
252 |
from a3 a21 have "(a,\<tau>)\<in>set \<Gamma>" by force |
|
253 |
with case1 a25 show False by force |
|
| 18106 | 254 |
qed |
255 |
qed |
|
| 18378 | 256 |
next |
257 |
assume case2: "a\<noteq>c" |
|
258 |
with a21 have a26: "(a,\<tau>)\<in>set \<Gamma>" by force |
|
259 |
from a23 a26 show "\<Gamma> \<turnstile> Var a:\<tau>" by force |
|
| 18106 | 260 |
qed |
261 |
next |
|
| 18313 | 262 |
case (App s1 s2) |
| 18383 | 263 |
have ih_s1: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s1[c::=t2]:\<tau>" by fact |
264 |
have ih_s2: "\<And>c \<sigma> \<tau> t2 \<Gamma>. ((c,\<sigma>)#\<Gamma>) \<turnstile> s2:\<tau> \<Longrightarrow> \<Gamma>\<turnstile> t2: \<sigma> \<Longrightarrow> \<Gamma> \<turnstile> s2[c::=t2]:\<tau>" by fact |
|
265 |
have "((c,\<sigma>)#\<Gamma>)\<turnstile>App s1 s2 : \<tau>" by fact |
|
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
266 |
hence "\<exists>\<tau>'. (c,\<sigma>)#\<Gamma> \<turnstile>s1:\<tau>'\<rightarrow>\<tau> \<and> (c,\<sigma>)#\<Gamma> \<turnstile> s2 : \<tau>'" by (rule t2_elim) |
| 18383 | 267 |
then obtain \<tau>' where "((c,\<sigma>)#\<Gamma>)\<turnstile>s1:\<tau>'\<rightarrow>\<tau>" and "((c,\<sigma>)#\<Gamma>)\<turnstile>s2:\<tau>'" by blast |
268 |
moreover |
|
269 |
have "\<Gamma> \<turnstile>t2:\<sigma>" by fact |
|
270 |
ultimately show "\<Gamma> \<turnstile> (App s1 s2)[c::=t2] : \<tau>" using ih_s1 ih_s2 by (simp, blast) |
|
| 18106 | 271 |
next |
| 18313 | 272 |
case (Lam a s) |
273 |
have "a\<sharp>\<Gamma>" "a\<sharp>\<sigma>" "a\<sharp>\<tau>" "a\<sharp>c" "a\<sharp>t2" by fact |
|
| 18106 | 274 |
hence f1: "a\<sharp>\<Gamma>" and f2: "a\<noteq>c" and f2': "c\<sharp>a" and f3: "a\<sharp>t2" and f4: "a\<sharp>((c,\<sigma>)#\<Gamma>)" |
275 |
by (auto simp add: fresh_atm fresh_prod fresh_list_cons) |
|
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
276 |
have c1: "(c,\<sigma>)#\<Gamma> \<turnstile> Lam [a].s : \<tau>" by fact |
|
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
277 |
hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> (a,\<tau>1)#(c,\<sigma>)#\<Gamma> \<turnstile> s : \<tau>2" using f4 by (auto dest: t3_elim) |
| 18378 | 278 |
then obtain \<tau>1 \<tau>2 where c11: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and c12: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" by force |
279 |
from c12 have "valid ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>)" by (rule typing_valid) |
|
280 |
hence ca: "valid \<Gamma>" and cb: "a\<sharp>\<Gamma>" and cc: "c\<sharp>\<Gamma>" |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
281 |
by (auto simp add: fresh_list_cons) |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
282 |
from c12 have c14: "(c,\<sigma>)#(a,\<tau>1)#\<Gamma> \<turnstile> s : \<tau>2" |
| 18378 | 283 |
proof - |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
284 |
have c2: "(a,\<tau>1)#(c,\<sigma>)#\<Gamma> \<subseteq> (c,\<sigma>)#(a,\<tau>1)#\<Gamma>" by force |
| 18378 | 285 |
have c3: "valid ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" |
286 |
by (rule v2, rule v2, auto simp add: fresh_list_cons fresh_prod ca cb cc f2' fresh_ty) |
|
287 |
from c12 c2 c3 show ?thesis by (force intro: weakening) |
|
| 18106 | 288 |
qed |
| 18378 | 289 |
assume c8: "\<Gamma> \<turnstile> t2 : \<sigma>" |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
290 |
have c81: "(a,\<tau>1)#\<Gamma> \<turnstile> t2 :\<sigma>" |
| 18378 | 291 |
proof - |
|
22650
0c5b22076fb3
tuned the proof of lemma pt_list_set_fresh (as suggested by Randy Pollack) and tuned the syntax for sub_contexts
urbanc
parents:
22542
diff
changeset
|
292 |
have c82: "\<Gamma> \<subseteq> (a,\<tau>1)#\<Gamma>" by force |
| 18378 | 293 |
have c83: "valid ((a,\<tau>1)#\<Gamma>)" using f1 ca by force |
294 |
with c8 c82 c83 show ?thesis by (force intro: weakening) |
|
295 |
qed |
|
296 |
show "\<Gamma> \<turnstile> (Lam [a].s)[c::=t2] : \<tau>" |
|
297 |
using c11 prems c14 c81 f1 by force |
|
| 18106 | 298 |
qed |
299 |
||
| 18378 | 300 |
lemma subject: |
| 18106 | 301 |
assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2" |
| 18378 | 302 |
and b: "\<Gamma> \<turnstile> t1:\<tau>" |
303 |
shows "\<Gamma> \<turnstile> t2:\<tau>" |
|
304 |
using a b |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
305 |
proof (nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct) |
|
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
306 |
case (b1 s1 s2 t) --"App-case left" |
| 18383 | 307 |
have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1:\<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact |
| 18378 | 308 |
have "\<Gamma> \<turnstile> App s1 t : \<tau>" by fact |
309 |
hence "\<exists>\<sigma>. \<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> t : \<sigma>" by (rule t2_elim) |
|
| 18383 | 310 |
then obtain \<sigma> where "\<Gamma> \<turnstile> s1 : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> t : \<sigma>" by blast |
311 |
with ih show "\<Gamma> \<turnstile> App s2 t : \<tau>" by blast |
|
| 18106 | 312 |
next |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
313 |
case (b2 s1 s2 t) --"App-case right" |
| 18383 | 314 |
have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact |
| 18378 | 315 |
have "\<Gamma> \<turnstile> App t s1 : \<tau>" by fact |
316 |
hence "\<exists>\<sigma>. \<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s1 : \<sigma>" by (rule t2_elim) |
|
| 18383 | 317 |
then obtain \<sigma> where "\<Gamma> \<turnstile> t : \<sigma>\<rightarrow>\<tau>" and "\<Gamma> \<turnstile> s1 : \<sigma>" by blast |
318 |
with ih show "\<Gamma> \<turnstile> App t s2 : \<tau>" by blast |
|
| 18106 | 319 |
next |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
320 |
case (b3 s1 s2 a) --"Lam-case" |
| 18383 | 321 |
have fr: "a\<sharp>\<Gamma>" "a\<sharp>\<tau>" by fact |
322 |
have ih: "\<And>\<Gamma> \<tau>. \<Gamma> \<turnstile> s1 : \<tau> \<Longrightarrow> \<Gamma> \<turnstile> s2 : \<tau>" by fact |
|
| 18378 | 323 |
have "\<Gamma> \<turnstile> Lam [a].s1 : \<tau>" by fact |
| 18383 | 324 |
with fr have "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by (simp add: t3_elim) |
325 |
then obtain \<tau>1 \<tau>2 where "\<tau>=\<tau>1\<rightarrow>\<tau>2" and "((a,\<tau>1)#\<Gamma>) \<turnstile> s1 : \<tau>2" by blast |
|
326 |
with ih show "\<Gamma> \<turnstile> Lam [a].s2 : \<tau>" using fr by blast |
|
| 18106 | 327 |
next |
| 22540 | 328 |
case (b4 a s2 s1) --"Beta-redex" |
| 18383 | 329 |
have fr: "a\<sharp>\<Gamma>" by fact |
| 18378 | 330 |
have "\<Gamma> \<turnstile> App (Lam [a].s1) s2 : \<tau>" by fact |
| 18383 | 331 |
hence "\<exists>\<sigma>. (\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s2 : \<sigma>)" by (simp add: t2_elim) |
332 |
then obtain \<sigma> where a1: "\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau>" and a2: "\<Gamma> \<turnstile> s2 : \<sigma>" by blast |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
333 |
from a1 have "((a,\<sigma>)#\<Gamma>) \<turnstile> s1 : \<tau>" using fr by (auto dest!: t3_elim simp add: ty.inject) |
| 18383 | 334 |
with a2 show "\<Gamma> \<turnstile> s1[a::=s2] : \<tau>" by (simp add: ty_subs) |
| 18106 | 335 |
qed |
336 |
||
| 18378 | 337 |
lemma subject_automatic: |
| 18106 | 338 |
assumes a: "t1\<longrightarrow>\<^isub>\<beta>t2" |
| 18378 | 339 |
and b: "\<Gamma> \<turnstile> t1:\<tau>" |
340 |
shows "\<Gamma> \<turnstile> t2:\<tau>" |
|
341 |
using a b |
|
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
342 |
apply(nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: Beta.strong_induct) |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
343 |
apply(auto dest!: t2_elim t3_elim intro: ty_subs simp add: ty.inject) |
| 18106 | 344 |
done |
345 |
||
346 |
subsection {* some facts about beta *}
|
|
347 |
||
348 |
constdefs |
|
349 |
"NORMAL" :: "lam \<Rightarrow> bool" |
|
350 |
"NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^isub>\<beta> t')" |
|
351 |
||
| 18383 | 352 |
lemma NORMAL_Var: |
353 |
shows "NORMAL (Var a)" |
|
354 |
proof - |
|
355 |
{ assume "\<exists>t'. (Var a) \<longrightarrow>\<^isub>\<beta> t'"
|
|
356 |
then obtain t' where "(Var a) \<longrightarrow>\<^isub>\<beta> t'" by blast |
|
357 |
hence False by (cases, auto) |
|
358 |
} |
|
359 |
thus "NORMAL (Var a)" by (force simp add: NORMAL_def) |
|
360 |
qed |
|
361 |
||
| 18106 | 362 |
constdefs |
363 |
"SN" :: "lam \<Rightarrow> bool" |
|
| 22271 | 364 |
"SN t \<equiv> termi Beta t" |
| 18106 | 365 |
|
| 18383 | 366 |
lemma SN_preserved: "\<lbrakk>SN(t1);t1\<longrightarrow>\<^isub>\<beta> t2\<rbrakk>\<Longrightarrow>SN(t2)" |
| 18106 | 367 |
apply(simp add: SN_def) |
368 |
apply(drule_tac a="t2" in acc_downward) |
|
369 |
apply(auto) |
|
370 |
done |
|
371 |
||
| 18383 | 372 |
lemma SN_intro: "(\<forall>t2. t1\<longrightarrow>\<^isub>\<beta>t2 \<longrightarrow> SN(t2))\<Longrightarrow>SN(t1)" |
| 18106 | 373 |
apply(simp add: SN_def) |
374 |
apply(rule accI) |
|
375 |
apply(auto) |
|
376 |
done |
|
377 |
||
378 |
section {* Candidates *}
|
|
379 |
||
|
22440
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
urbanc
parents:
22420
diff
changeset
|
380 |
consts |
| 18106 | 381 |
RED :: "ty \<Rightarrow> lam set" |
|
22440
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
urbanc
parents:
22420
diff
changeset
|
382 |
|
|
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
urbanc
parents:
22420
diff
changeset
|
383 |
nominal_primrec |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
384 |
"RED (TVar X) = {t. SN(t)}"
|
|
22440
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
urbanc
parents:
22420
diff
changeset
|
385 |
"RED (\<tau>\<rightarrow>\<sigma>) = {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
|
|
7e4f4f19002f
deleted function for defining candidates and used nominal_primrec instead
urbanc
parents:
22420
diff
changeset
|
386 |
apply(rule TrueI)+ |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
387 |
done |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
388 |
|
| 18106 | 389 |
constdefs |
390 |
NEUT :: "lam \<Rightarrow> bool" |
|
391 |
"NEUT t \<equiv> (\<exists>a. t=Var a)\<or>(\<exists>t1 t2. t=App t1 t2)" |
|
392 |
||
393 |
(* a slight hack to get the first element of applications *) |
|
| 22271 | 394 |
inductive2 FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
|
395 |
where |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
396 |
fst[intro!]: "(App t s) \<guillemotright> t" |
| 18106 | 397 |
|
| 18378 | 398 |
lemma fst_elim[elim!]: |
399 |
shows "(App t s) \<guillemotright> t' \<Longrightarrow> t=t'" |
|
| 22271 | 400 |
apply(ind_cases2 "App t s \<guillemotright> t'") |
| 18106 | 401 |
apply(simp add: lam.inject) |
402 |
done |
|
403 |
||
404 |
lemma qq3: "SN(App t s)\<Longrightarrow>SN(t)" |
|
405 |
apply(simp add: SN_def) |
|
| 22271 | 406 |
apply(subgoal_tac "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> termi Beta z")(*A*) |
| 18106 | 407 |
apply(force) |
408 |
(*A*) |
|
409 |
apply(erule acc_induct) |
|
410 |
apply(clarify) |
|
| 22271 | 411 |
apply(ind_cases2 "x \<guillemotright> z" for x z) |
| 18106 | 412 |
apply(clarify) |
413 |
apply(rule accI) |
|
414 |
apply(auto intro: b1) |
|
415 |
done |
|
416 |
||
| 18383 | 417 |
section {* Candidates *}
|
418 |
||
| 18106 | 419 |
constdefs |
| 18383 | 420 |
"CR1" :: "ty \<Rightarrow> bool" |
421 |
"CR1 \<tau> \<equiv> \<forall> t. (t\<in>RED \<tau> \<longrightarrow> SN(t))" |
|
| 18106 | 422 |
|
| 18383 | 423 |
"CR2" :: "ty \<Rightarrow> bool" |
424 |
"CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^isub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>" |
|
| 18106 | 425 |
|
| 18383 | 426 |
"CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool" |
427 |
"CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^isub>\<beta> t' \<longrightarrow> t'\<in>RED \<tau>" |
|
| 18106 | 428 |
|
| 18383 | 429 |
"CR3" :: "ty \<Rightarrow> bool" |
430 |
"CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>" |
|
| 18106 | 431 |
|
| 18383 | 432 |
"CR4" :: "ty \<Rightarrow> bool" |
433 |
"CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>" |
|
| 18106 | 434 |
|
435 |
lemma CR3_CR4: "CR3 \<tau> \<Longrightarrow> CR4 \<tau>" |
|
436 |
apply(simp (no_asm_use) add: CR3_def CR3_RED_def CR4_def NORMAL_def) |
|
437 |
apply(blast) |
|
438 |
done |
|
439 |
||
440 |
lemma sub_ind: |
|
441 |
"SN(u)\<Longrightarrow>(u\<in>RED \<tau>\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>\<and>CR3 \<sigma>\<and>CR3_RED t (\<tau>\<rightarrow>\<sigma>))\<longrightarrow>(App t u)\<in>RED \<sigma>))" |
|
442 |
apply(simp add: SN_def) |
|
443 |
apply(erule acc_induct) |
|
444 |
apply(auto) |
|
445 |
apply(simp add: CR3_def) |
|
446 |
apply(rotate_tac 5) |
|
447 |
apply(drule_tac x="App t x" in spec) |
|
448 |
apply(drule mp) |
|
449 |
apply(rule conjI) |
|
450 |
apply(force simp only: NEUT_def) |
|
451 |
apply(simp (no_asm) add: CR3_RED_def) |
|
452 |
apply(clarify) |
|
| 22271 | 453 |
apply(ind_cases2 "App t x \<longrightarrow>\<^isub>\<beta> t'" for x t t') |
| 18106 | 454 |
apply(simp_all add: lam.inject) |
455 |
apply(simp only: CR3_RED_def) |
|
456 |
apply(drule_tac x="s2" in spec) |
|
457 |
apply(simp) |
|
458 |
apply(drule_tac x="s2" in spec) |
|
459 |
apply(simp) |
|
460 |
apply(drule mp) |
|
461 |
apply(simp (no_asm_use) add: CR2_def) |
|
462 |
apply(blast) |
|
463 |
apply(drule_tac x="ta" in spec) |
|
464 |
apply(force) |
|
465 |
apply(auto simp only: NEUT_def lam.inject lam.distinct) |
|
466 |
done |
|
467 |
||
| 18383 | 468 |
lemma RED_props: |
469 |
shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>" |
|
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
470 |
proof (nominal_induct \<tau> rule: ty.induct) |
| 18611 | 471 |
case (TVar a) |
472 |
{ case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
|
|
473 |
next |
|
474 |
case 2 show "CR2 (TVar a)" by (force intro: SN_preserved simp add: CR2_def) |
|
475 |
next |
|
476 |
case 3 show "CR3 (TVar a)" by (force intro: SN_intro simp add: CR3_def CR3_RED_def) |
|
477 |
} |
|
|
18599
e01112713fdc
changed PRO_RED proof to conform with the new induction rules
urbanc
parents:
18383
diff
changeset
|
478 |
next |
| 18611 | 479 |
case (TArr \<tau>1 \<tau>2) |
480 |
{ case 1
|
|
481 |
have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact |
|
482 |
have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact |
|
483 |
show "CR1 (\<tau>1 \<rightarrow> \<tau>2)" |
|
484 |
proof (simp add: CR1_def, intro strip) |
|
485 |
fix t |
|
486 |
assume a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2" |
|
487 |
from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_CR4) |
|
488 |
moreover |
|
489 |
have "NEUT (Var a)" by (force simp add: NEUT_def) |
|
490 |
moreover |
|
491 |
have "NORMAL (Var a)" by (rule NORMAL_Var) |
|
492 |
ultimately have "(Var a)\<in> RED \<tau>1" by (simp add: CR4_def) |
|
493 |
with a have "App t (Var a) \<in> RED \<tau>2" by simp |
|
494 |
hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def) |
|
495 |
thus "SN(t)" by (rule qq3) |
|
496 |
qed |
|
497 |
next |
|
498 |
case 2 |
|
499 |
have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact |
|
500 |
have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact |
|
501 |
show "CR2 (\<tau>1 \<rightarrow> \<tau>2)" |
|
502 |
proof (simp add: CR2_def, intro strip) |
|
503 |
fix t1 t2 u |
|
504 |
assume "(\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t1 u \<in> RED \<tau>2) \<and> t1 \<longrightarrow>\<^isub>\<beta> t2" |
|
505 |
and "u \<in> RED \<tau>1" |
|
506 |
hence "t1 \<longrightarrow>\<^isub>\<beta> t2" and "App t1 u \<in> RED \<tau>2" by simp_all |
|
507 |
thus "App t2 u \<in> RED \<tau>2" using ih_CR2_\<tau>2 by (force simp add: CR2_def) |
|
508 |
qed |
|
509 |
next |
|
510 |
case 3 |
|
511 |
have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact |
|
512 |
have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact |
|
513 |
have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact |
|
514 |
show "CR3 (\<tau>1 \<rightarrow> \<tau>2)" |
|
515 |
proof (simp add: CR3_def, intro strip) |
|
516 |
fix t u |
|
517 |
assume a1: "u \<in> RED \<tau>1" |
|
518 |
assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)" |
|
519 |
from a1 have "SN(u)" using ih_CR1_\<tau>1 by (simp add: CR1_def) |
|
520 |
hence "u\<in>RED \<tau>1\<longrightarrow>(\<forall>t. (NEUT t\<and>CR2 \<tau>1\<and>CR3 \<tau>2\<and>CR3_RED t (\<tau>1\<rightarrow>\<tau>2))\<longrightarrow>(App t u)\<in>RED \<tau>2)" |
|
521 |
by (rule sub_ind) |
|
522 |
with a1 a2 show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 by simp |
|
523 |
qed |
|
524 |
} |
|
| 18383 | 525 |
qed |
526 |
||
| 18106 | 527 |
lemma double_acc_aux: |
| 22271 | 528 |
assumes a_acc: "acc r a" |
529 |
and b_acc: "acc r b" |
|
| 18106 | 530 |
and hyp: "\<And>x z. |
| 22271 | 531 |
(\<And>y. r y x \<Longrightarrow> acc r y) \<Longrightarrow> |
532 |
(\<And>y. r y x \<Longrightarrow> P y z) \<Longrightarrow> |
|
533 |
(\<And>u. r u z \<Longrightarrow> acc r u) \<Longrightarrow> |
|
534 |
(\<And>u. r u z \<Longrightarrow> P x u) \<Longrightarrow> P x z" |
|
| 18106 | 535 |
shows "P a b" |
536 |
proof - |
|
537 |
from a_acc |
|
| 22271 | 538 |
have r: "\<And>b. acc r b \<Longrightarrow> P a b" |
| 18106 | 539 |
proof (induct a rule: acc.induct) |
540 |
case (accI x) |
|
541 |
note accI' = accI |
|
| 22271 | 542 |
have "acc r b" . |
| 18106 | 543 |
thus ?case |
544 |
proof (induct b rule: acc.induct) |
|
545 |
case (accI y) |
|
546 |
show ?case |
|
547 |
apply (rule hyp) |
|
548 |
apply (erule accI') |
|
549 |
apply (erule accI') |
|
550 |
apply (rule acc.accI) |
|
551 |
apply (erule accI) |
|
552 |
apply (erule accI) |
|
553 |
apply (erule accI) |
|
554 |
done |
|
555 |
qed |
|
556 |
qed |
|
557 |
from b_acc show ?thesis by (rule r) |
|
558 |
qed |
|
559 |
||
560 |
lemma double_acc: |
|
| 22271 | 561 |
"\<lbrakk>acc r a; acc r b; \<forall>x z. ((\<forall>y. r y x \<longrightarrow> P y z) \<and> (\<forall>u. r u z \<longrightarrow> P x u)) \<longrightarrow> P x z\<rbrakk> \<Longrightarrow> P a b" |
| 18106 | 562 |
apply(rule_tac r="r" in double_acc_aux) |
563 |
apply(assumption)+ |
|
564 |
apply(blast) |
|
565 |
done |
|
566 |
||
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
567 |
lemma abs_RED: "(\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>)\<longrightarrow>Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)" |
| 18106 | 568 |
apply(simp) |
569 |
apply(clarify) |
|
| 22271 | 570 |
apply(subgoal_tac "termi Beta t")(*1*) |
| 18106 | 571 |
apply(erule rev_mp) |
572 |
apply(subgoal_tac "u \<in> RED \<tau>")(*A*) |
|
573 |
apply(erule rev_mp) |
|
574 |
apply(rule_tac a="t" and b="u" in double_acc) |
|
575 |
apply(assumption) |
|
576 |
apply(subgoal_tac "CR1 \<tau>")(*A*) |
|
577 |
apply(simp add: CR1_def SN_def) |
|
578 |
(*A*) |
|
579 |
apply(force simp add: RED_props) |
|
580 |
apply(simp) |
|
581 |
apply(clarify) |
|
582 |
apply(subgoal_tac "CR3 \<sigma>")(*B*) |
|
583 |
apply(simp add: CR3_def) |
|
584 |
apply(rotate_tac 6) |
|
585 |
apply(drule_tac x="App(Lam[x].xa ) z" in spec) |
|
586 |
apply(drule mp) |
|
587 |
apply(rule conjI) |
|
588 |
apply(force simp add: NEUT_def) |
|
589 |
apply(simp add: CR3_RED_def) |
|
590 |
apply(clarify) |
|
| 22271 | 591 |
apply(ind_cases2 "App(Lam[x].xa) z \<longrightarrow>\<^isub>\<beta> t'" for xa z t') |
| 18106 | 592 |
apply(auto simp add: lam.inject lam.distinct) |
593 |
apply(drule beta_abs) |
|
594 |
apply(auto) |
|
595 |
apply(drule_tac x="t''" in spec) |
|
596 |
apply(simp) |
|
597 |
apply(drule mp) |
|
598 |
apply(clarify) |
|
599 |
apply(drule_tac x="s" in bspec) |
|
600 |
apply(assumption) |
|
601 |
apply(subgoal_tac "xa [ x ::= s ] \<longrightarrow>\<^isub>\<beta> t'' [ x ::= s ]")(*B*) |
|
602 |
apply(subgoal_tac "CR2 \<sigma>")(*C*) |
|
603 |
apply(simp (no_asm_use) add: CR2_def) |
|
604 |
apply(blast) |
|
605 |
(*C*) |
|
606 |
apply(force simp add: RED_props) |
|
607 |
(*B*) |
|
608 |
apply(force intro!: beta_subst) |
|
609 |
apply(assumption) |
|
610 |
apply(rotate_tac 3) |
|
611 |
apply(drule_tac x="s2" in spec) |
|
612 |
apply(subgoal_tac "s2\<in>RED \<tau>")(*D*) |
|
613 |
apply(simp) |
|
614 |
(*D*) |
|
615 |
apply(subgoal_tac "CR2 \<tau>")(*E*) |
|
616 |
apply(simp (no_asm_use) add: CR2_def) |
|
617 |
apply(blast) |
|
618 |
(*E*) |
|
619 |
apply(force simp add: RED_props) |
|
620 |
apply(simp add: alpha) |
|
621 |
apply(erule disjE) |
|
622 |
apply(force) |
|
623 |
apply(auto) |
|
624 |
apply(simp add: subst_rename) |
|
625 |
apply(drule_tac x="z" in bspec) |
|
626 |
apply(assumption) |
|
627 |
(*B*) |
|
628 |
apply(force simp add: RED_props) |
|
629 |
(*1*) |
|
630 |
apply(drule_tac x="Var x" in bspec) |
|
631 |
apply(subgoal_tac "CR3 \<tau>")(*2*) |
|
632 |
apply(drule CR3_CR4) |
|
633 |
apply(simp add: CR4_def) |
|
634 |
apply(drule_tac x="Var x" in spec) |
|
635 |
apply(drule mp) |
|
636 |
apply(rule conjI) |
|
637 |
apply(force simp add: NEUT_def) |
|
638 |
apply(simp add: NORMAL_def) |
|
639 |
apply(clarify) |
|
| 22271 | 640 |
apply(ind_cases2 "Var x \<longrightarrow>\<^isub>\<beta> t'" for t') |
| 18106 | 641 |
apply(auto simp add: lam.inject lam.distinct) |
642 |
apply(force simp add: RED_props) |
|
643 |
apply(simp add: id_subs) |
|
644 |
apply(subgoal_tac "CR1 \<sigma>")(*3*) |
|
645 |
apply(simp add: CR1_def SN_def) |
|
646 |
(*3*) |
|
647 |
apply(force simp add: RED_props) |
|
648 |
done |
|
649 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
650 |
lemma psubst_subst: |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
651 |
assumes h:"c\<sharp>\<theta>" |
| 22420 | 652 |
shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>" |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
653 |
using h |
|
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
654 |
by (nominal_induct t avoiding: \<theta> c s rule: lam.induct) |
| 22420 | 655 |
(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') |
656 |
||
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
657 |
|
| 22420 | 658 |
abbreviation |
659 |
mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" ("_ maps _ to _" [55,55,55] 55)
|
|
660 |
where |
|
661 |
"\<theta> maps x to e\<equiv> (lookup \<theta> x) = e" |
|
662 |
||
663 |
abbreviation |
|
664 |
closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" ("_ closes _" [55,55] 55)
|
|
665 |
where |
|
666 |
"\<theta> closes \<Gamma> \<equiv> \<forall>x T. ((x,T) \<in> set \<Gamma> \<longrightarrow> (\<exists>t. \<theta> maps x to t \<and> t \<in> RED T))" |
|
|
21107
e69c0e82955a
new file for defining functions in the lambda-calculus
urbanc
parents:
19972
diff
changeset
|
667 |
|
| 18106 | 668 |
lemma all_RED: |
| 22420 | 669 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
670 |
and b: "\<theta> closes \<Gamma>" |
|
671 |
shows "\<theta><t> \<in> RED \<tau>" |
|
| 18345 | 672 |
using a b |
|
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18654
diff
changeset
|
673 |
proof(nominal_induct t avoiding: \<Gamma> \<tau> \<theta> rule: lam.induct) |
| 18345 | 674 |
case (Lam a t) --"lambda case" |
| 22420 | 675 |
have ih: "\<And>\<Gamma> \<tau> \<theta>. \<Gamma> \<turnstile> t : \<tau> \<Longrightarrow> \<theta> closes \<Gamma> \<Longrightarrow> \<theta><t> \<in> RED \<tau>" |
676 |
and \<theta>_cond: "\<theta> closes \<Gamma>" |
|
| 18345 | 677 |
and fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" |
678 |
and "\<Gamma> \<turnstile> Lam [a].t:\<tau>" by fact |
|
679 |
hence "\<exists>\<tau>1 \<tau>2. \<tau>=\<tau>1\<rightarrow>\<tau>2 \<and> ((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" using t3_elim fresh by simp |
|
680 |
then obtain \<tau>1 \<tau>2 where \<tau>_inst: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and typing: "((a,\<tau>1)#\<Gamma>)\<turnstile>t:\<tau>2" by blast |
|
| 22420 | 681 |
from ih have "\<forall>s\<in>RED \<tau>1. \<theta><t>[a::=s] \<in> RED \<tau>2" using fresh typing \<theta>_cond |
| 18345 | 682 |
by (force dest: fresh_context simp add: psubst_subst) |
| 22420 | 683 |
hence "(Lam [a].(\<theta><t>)) \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" by (simp only: abs_RED) |
684 |
thus "\<theta><(Lam [a].t)> \<in> RED \<tau>" using fresh \<tau>_inst by simp |
|
| 18345 | 685 |
qed (force dest!: t1_elim t2_elim)+ |
686 |
||
| 18383 | 687 |
(* identity substitution generated from a context \<Gamma> *) |
| 18382 | 688 |
consts |
689 |
"id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list" |
|
690 |
primrec |
|
691 |
"id [] = []" |
|
692 |
"id (x#\<Gamma>) = ((fst x),Var (fst x))#(id \<Gamma>)" |
|
693 |
||
694 |
lemma id_var: |
|
| 22420 | 695 |
shows "(id \<Gamma>) maps a to Var a" |
| 18382 | 696 |
apply(induct \<Gamma>, auto) |
697 |
done |
|
698 |
||
699 |
lemma id_fresh: |
|
700 |
fixes a::"name" |
|
701 |
assumes a: "a\<sharp>\<Gamma>" |
|
702 |
shows "a\<sharp>(id \<Gamma>)" |
|
703 |
using a |
|
704 |
apply(induct \<Gamma>) |
|
705 |
apply(auto simp add: fresh_list_nil fresh_list_cons fresh_prod) |
|
706 |
done |
|
707 |
||
708 |
lemma id_apply: |
|
| 22420 | 709 |
shows "(id \<Gamma>)<t> = t" |
|
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18654
diff
changeset
|
710 |
apply(nominal_induct t avoiding: \<Gamma> rule: lam.induct) |
| 18382 | 711 |
apply(auto) |
| 22420 | 712 |
apply(rule id_var) |
| 18382 | 713 |
apply(drule id_fresh)+ |
714 |
apply(simp) |
|
715 |
done |
|
716 |
||
717 |
lemma id_mem: |
|
718 |
assumes a: "(a,\<tau>)\<in>set \<Gamma>" |
|
| 22420 | 719 |
shows "lookup (id \<Gamma>) a = Var a" |
| 18382 | 720 |
using a |
721 |
apply(induct \<Gamma>, auto) |
|
722 |
done |
|
723 |
||
| 18383 | 724 |
lemma id_prop: |
| 22420 | 725 |
shows "(id \<Gamma>) closes \<Gamma>" |
| 18383 | 726 |
apply(auto) |
| 18382 | 727 |
apply(simp add: id_mem) |
| 22420 | 728 |
apply(subgoal_tac "CR3 T") --"A" |
| 18382 | 729 |
apply(drule CR3_CR4) |
730 |
apply(simp add: CR4_def) |
|
| 22420 | 731 |
apply(drule_tac x="Var x" in spec) |
| 18383 | 732 |
apply(force simp add: NEUT_def NORMAL_Var) |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
733 |
--"A" |
| 18383 | 734 |
apply(rule RED_props) |
| 18382 | 735 |
done |
736 |
||
| 18383 | 737 |
lemma typing_implies_RED: |
738 |
assumes a: "\<Gamma>\<turnstile>t:\<tau>" |
|
739 |
shows "t \<in> RED \<tau>" |
|
740 |
proof - |
|
| 22420 | 741 |
have "(id \<Gamma>)<t>\<in>RED \<tau>" |
| 18383 | 742 |
proof - |
| 22420 | 743 |
have "(id \<Gamma>) closes \<Gamma>" by (rule id_prop) |
| 18383 | 744 |
with a show ?thesis by (rule all_RED) |
745 |
qed |
|
746 |
thus"t \<in> RED \<tau>" by (simp add: id_apply) |
|
747 |
qed |
|
748 |
||
749 |
lemma typing_implies_SN: |
|
750 |
assumes a: "\<Gamma>\<turnstile>t:\<tau>" |
|
751 |
shows "SN(t)" |
|
752 |
proof - |
|
753 |
from a have "t \<in> RED \<tau>" by (rule typing_implies_RED) |
|
754 |
moreover |
|
755 |
have "CR1 \<tau>" by (rule RED_props) |
|
756 |
ultimately show "SN(t)" by (simp add: CR1_def) |
|
757 |
qed |
|
| 18382 | 758 |
|
759 |
end |