author | wenzelm |
Wed, 25 Jun 2025 16:35:25 +0200 | |
changeset 82768 | 8f866fd6fae1 |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
19496 | 1 |
theory SN |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
2 |
imports Lam_Funs |
18106 | 3 |
begin |
4 |
||
63167 | 5 |
text \<open>Strong Normalisation proof from the Proofs and Types book\<close> |
18106 | 6 |
|
63167 | 7 |
section \<open>Beta Reduction\<close> |
18106 | 8 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
9 |
lemma subst_rename: |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
10 |
assumes a: "c\<sharp>t1" |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
11 |
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
|
12 |
using a |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
13 |
by (nominal_induct t1 avoiding: a c t2 rule: lam.strong_induct) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
14 |
(auto simp add: calc_atm fresh_atm abs_fresh) |
18106 | 15 |
|
18313 | 16 |
lemma forget: |
17 |
assumes a: "a\<sharp>t1" |
|
18 |
shows "t1[a::=t2] = t1" |
|
19 |
using a |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
20 |
by (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
21 |
(auto simp add: abs_fresh fresh_atm) |
18106 | 22 |
|
18313 | 23 |
lemma fresh_fact: |
24 |
fixes a::"name" |
|
23970 | 25 |
assumes a: "a\<sharp>t1" "a\<sharp>t2" |
22540 | 26 |
shows "a\<sharp>t1[b::=t2]" |
23970 | 27 |
using a |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
28 |
by (nominal_induct t1 avoiding: a b t2 rule: lam.strong_induct) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
29 |
(auto simp add: abs_fresh fresh_atm) |
18106 | 30 |
|
22540 | 31 |
lemma fresh_fact': |
32 |
fixes a::"name" |
|
33 |
assumes a: "a\<sharp>t2" |
|
34 |
shows "a\<sharp>t1[a::=t2]" |
|
35 |
using a |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
36 |
by (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) |
22540 | 37 |
(auto simp add: abs_fresh fresh_atm) |
38 |
||
18383 | 39 |
lemma subst_lemma: |
18313 | 40 |
assumes a: "x\<noteq>y" |
41 |
and b: "x\<sharp>L" |
|
42 |
shows "M[x::=N][y::=L] = M[y::=L][x::=N[y::=L]]" |
|
43 |
using a b |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
44 |
by (nominal_induct M avoiding: x y N L rule: lam.strong_induct) |
18313 | 45 |
(auto simp add: fresh_fact forget) |
18106 | 46 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
47 |
lemma id_subs: |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
48 |
shows "t[x::=Var x] = t" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
49 |
by (nominal_induct t avoiding: x rule: lam.strong_induct) |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
50 |
(simp_all add: fresh_atm) |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
51 |
|
26772 | 52 |
lemma lookup_fresh: |
53 |
fixes z::"name" |
|
54 |
assumes "z\<sharp>\<theta>" "z\<sharp>x" |
|
55 |
shows "z\<sharp> lookup \<theta> x" |
|
56 |
using assms |
|
57 |
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons) |
|
58 |
||
59 |
lemma lookup_fresh': |
|
60 |
assumes "z\<sharp>\<theta>" |
|
61 |
shows "lookup \<theta> z = Var z" |
|
62 |
using assms |
|
63 |
by (induct rule: lookup.induct) |
|
64 |
(auto simp add: fresh_list_cons fresh_prod fresh_atm) |
|
65 |
||
23142 | 66 |
lemma psubst_subst: |
67 |
assumes h:"c\<sharp>\<theta>" |
|
68 |
shows "(\<theta><t>)[c::=s] = ((c,s)#\<theta>)<t>" |
|
69 |
using h |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
70 |
by (nominal_induct t avoiding: \<theta> c s rule: lam.strong_induct) |
23142 | 71 |
(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh') |
72 |
||
23760 | 73 |
inductive |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
74 |
Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<longrightarrow>\<^sub>\<beta> _\<close> [80,80] 80) |
22271 | 75 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
76 |
b1[intro!]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App s1 t \<longrightarrow>\<^sub>\<beta> App s2 t" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
77 |
| b2[intro!]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
78 |
| b3[intro!]: "s1\<longrightarrow>\<^sub>\<beta>s2 \<Longrightarrow> Lam [a].s1 \<longrightarrow>\<^sub>\<beta> Lam [a].s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
79 |
| b4[intro!]: "a\<sharp>s2 \<Longrightarrow> App (Lam [a].s1) s2\<longrightarrow>\<^sub>\<beta> (s1[a::=s2])" |
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
80 |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
81 |
equivariance Beta |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
82 |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
83 |
nominal_inductive Beta |
22540 | 84 |
by (simp_all add: abs_fresh fresh_fact') |
18106 | 85 |
|
25831 | 86 |
lemma beta_preserves_fresh: |
87 |
fixes a::"name" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
88 |
assumes a: "t\<longrightarrow>\<^sub>\<beta> s" |
25831 | 89 |
shows "a\<sharp>t \<Longrightarrow> a\<sharp>s" |
18378 | 90 |
using a |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
91 |
by (nominal_induct t s avoiding: a rule: Beta.strong_induct) |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
92 |
(auto simp add: abs_fresh fresh_fact fresh_atm) |
18378 | 93 |
|
23970 | 94 |
lemma beta_abs: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
95 |
assumes a: "Lam [a].t\<longrightarrow>\<^sub>\<beta> t'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
96 |
shows "\<exists>t''. t'=Lam [a].t'' \<and> t\<longrightarrow>\<^sub>\<beta> t''" |
25831 | 97 |
proof - |
98 |
have "a\<sharp>Lam [a].t" by (simp add: abs_fresh) |
|
99 |
with a have "a\<sharp>t'" by (simp add: beta_preserves_fresh) |
|
100 |
with a show ?thesis |
|
101 |
by (cases rule: Beta.strong_cases[where a="a" and aa="a"]) |
|
102 |
(auto simp add: lam.inject abs_fresh alpha) |
|
103 |
qed |
|
18106 | 104 |
|
18313 | 105 |
lemma beta_subst: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
106 |
assumes a: "M \<longrightarrow>\<^sub>\<beta> M'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
107 |
shows "M[x::=N]\<longrightarrow>\<^sub>\<beta> M'[x::=N]" |
18106 | 108 |
using a |
23142 | 109 |
by (nominal_induct M M' avoiding: x N rule: Beta.strong_induct) |
110 |
(auto simp add: fresh_atm subst_lemma fresh_fact) |
|
18106 | 111 |
|
63167 | 112 |
section \<open>types\<close> |
18383 | 113 |
|
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
114 |
nominal_datatype ty = |
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
115 |
TVar "nat" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
116 |
| TArr "ty" "ty" (infix \<open>\<rightarrow>\<close> 200) |
18106 | 117 |
|
118 |
lemma fresh_ty: |
|
119 |
fixes a ::"name" |
|
120 |
and \<tau> ::"ty" |
|
121 |
shows "a\<sharp>\<tau>" |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
122 |
by (nominal_induct \<tau> rule: ty.strong_induct) |
25867 | 123 |
(auto simp add: fresh_nat) |
18106 | 124 |
|
125 |
(* valid contexts *) |
|
126 |
||
23760 | 127 |
inductive |
23142 | 128 |
valid :: "(name\<times>ty) list \<Rightarrow> bool" |
22271 | 129 |
where |
130 |
v1[intro]: "valid []" |
|
131 |
| v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)" |
|
18106 | 132 |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
133 |
equivariance valid |
18106 | 134 |
|
135 |
(* typing judgements *) |
|
136 |
||
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
137 |
lemma fresh_context: |
18106 | 138 |
fixes \<Gamma> :: "(name\<times>ty)list" |
139 |
and a :: "name" |
|
18378 | 140 |
assumes a: "a\<sharp>\<Gamma>" |
141 |
shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)" |
|
142 |
using a |
|
23970 | 143 |
by (induct \<Gamma>) |
144 |
(auto simp add: fresh_prod fresh_list_cons fresh_atm) |
|
18106 | 145 |
|
23760 | 146 |
inductive |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
147 |
typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (\<open>_ \<turnstile> _ : _\<close> [60,60,60] 60) |
22271 | 148 |
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
|
149 |
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
|
150 |
| t2[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>" |
22271 | 151 |
| t3[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>" |
18106 | 152 |
|
22730
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
153 |
equivariance typing |
8bcc8809ed3b
nominal_inductive no longer proves equivariance.
berghofe
parents:
22650
diff
changeset
|
154 |
|
22538
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
155 |
nominal_inductive typing |
3ccb92dfb5e9
tuned proofs (taking full advantage of nominal_inductive)
urbanc
parents:
22531
diff
changeset
|
156 |
by (simp_all add: abs_fresh fresh_ty) |
18106 | 157 |
|
63167 | 158 |
subsection \<open>a fact about beta\<close> |
18106 | 159 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
160 |
definition "NORMAL" :: "lam \<Rightarrow> bool" where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
161 |
"NORMAL t \<equiv> \<not>(\<exists>t'. t\<longrightarrow>\<^sub>\<beta> t')" |
18106 | 162 |
|
18383 | 163 |
lemma NORMAL_Var: |
164 |
shows "NORMAL (Var a)" |
|
165 |
proof - |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
166 |
{ assume "\<exists>t'. (Var a) \<longrightarrow>\<^sub>\<beta> t'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
167 |
then obtain t' where "(Var a) \<longrightarrow>\<^sub>\<beta> t'" by blast |
25867 | 168 |
hence False by (cases) (auto) |
18383 | 169 |
} |
25867 | 170 |
thus "NORMAL (Var a)" by (auto simp add: NORMAL_def) |
18383 | 171 |
qed |
172 |
||
63167 | 173 |
text \<open>Inductive version of Strong Normalisation\<close> |
23970 | 174 |
inductive |
175 |
SN :: "lam \<Rightarrow> bool" |
|
176 |
where |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
177 |
SN_intro: "(\<And>t'. t \<longrightarrow>\<^sub>\<beta> t' \<Longrightarrow> SN t') \<Longrightarrow> SN t" |
23970 | 178 |
|
179 |
lemma SN_preserved: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
180 |
assumes a: "SN t1" "t1\<longrightarrow>\<^sub>\<beta> t2" |
23970 | 181 |
shows "SN t2" |
182 |
using a |
|
183 |
by (cases) (auto) |
|
18106 | 184 |
|
23970 | 185 |
lemma double_SN_aux: |
186 |
assumes a: "SN a" |
|
187 |
and b: "SN b" |
|
188 |
and hyp: "\<And>x z. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
189 |
\<lbrakk>\<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> SN y; \<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> P y z; |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
190 |
\<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> SN u; \<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z" |
23970 | 191 |
shows "P a b" |
192 |
proof - |
|
193 |
from a |
|
194 |
have r: "\<And>b. SN b \<Longrightarrow> P a b" |
|
195 |
proof (induct a rule: SN.SN.induct) |
|
196 |
case (SN_intro x) |
|
197 |
note SNI' = SN_intro |
|
198 |
have "SN b" by fact |
|
199 |
thus ?case |
|
200 |
proof (induct b rule: SN.SN.induct) |
|
201 |
case (SN_intro y) |
|
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
202 |
with SNI' show ?case |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
203 |
by (metis SN.simps hyp) |
23970 | 204 |
qed |
205 |
qed |
|
206 |
from b show ?thesis by (rule r) |
|
207 |
qed |
|
18106 | 208 |
|
23970 | 209 |
lemma double_SN[consumes 2]: |
210 |
assumes a: "SN a" |
|
211 |
and b: "SN b" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
212 |
and c: "\<And>x z. \<lbrakk>\<And>y. x \<longrightarrow>\<^sub>\<beta> y \<Longrightarrow> P y z; \<And>u. z \<longrightarrow>\<^sub>\<beta> u \<Longrightarrow> P x u\<rbrakk> \<Longrightarrow> P x z" |
23970 | 213 |
shows "P a b" |
214 |
using a b c |
|
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
215 |
by (smt (verit, best) double_SN_aux) |
18106 | 216 |
|
63167 | 217 |
section \<open>Candidates\<close> |
18106 | 218 |
|
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
219 |
nominal_primrec |
18106 | 220 |
RED :: "ty \<Rightarrow> lam set" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
221 |
where |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
222 |
"RED (TVar X) = {t. SN(t)}" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
223 |
| "RED (\<tau>\<rightarrow>\<sigma>) = {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}" |
23970 | 224 |
by (rule TrueI)+ |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
225 |
|
63167 | 226 |
text \<open>neutral terms\<close> |
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
227 |
definition NEUT :: "lam \<Rightarrow> bool" where |
23970 | 228 |
"NEUT t \<equiv> (\<exists>a. t = Var a) \<or> (\<exists>t1 t2. t = App t1 t2)" |
18106 | 229 |
|
230 |
(* a slight hack to get the first element of applications *) |
|
23970 | 231 |
(* this is needed to get (SN t) from SN (App t s) *) |
23760 | 232 |
inductive |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
233 |
FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (\<open> _ \<guillemotright> _\<close> [80,80] 80) |
22271 | 234 |
where |
22418
49e2d9744ae1
major update of the nominal package; there is now an infrastructure
urbanc
parents:
22271
diff
changeset
|
235 |
fst[intro!]: "(App t s) \<guillemotright> t" |
18106 | 236 |
|
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
237 |
nominal_primrec |
24899 | 238 |
fst_app_aux::"lam\<Rightarrow>lam option" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
239 |
where |
24899 | 240 |
"fst_app_aux (Var a) = None" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
241 |
| "fst_app_aux (App t1 t2) = Some t1" |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27272
diff
changeset
|
242 |
| "fst_app_aux (Lam [x].t) = None" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
243 |
by (finite_guess | simp add: fresh_none | fresh_guess)+ |
24899 | 244 |
|
245 |
definition |
|
246 |
fst_app_def[simp]: "fst_app t = the (fst_app_aux t)" |
|
247 |
||
23970 | 248 |
lemma SN_of_FST_of_App: |
249 |
assumes a: "SN (App t s)" |
|
24899 | 250 |
shows "SN (fst_app (App t s))" |
23970 | 251 |
using a |
252 |
proof - |
|
253 |
from a have "\<forall>z. (App t s \<guillemotright> z) \<longrightarrow> SN z" |
|
254 |
by (induct rule: SN.SN.induct) |
|
255 |
(blast elim: FST.cases intro: SN_intro) |
|
24899 | 256 |
then have "SN t" by blast |
257 |
then show "SN (fst_app (App t s))" by simp |
|
23970 | 258 |
qed |
18106 | 259 |
|
63167 | 260 |
section \<open>Candidates\<close> |
18383 | 261 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
262 |
definition "CR1" :: "ty \<Rightarrow> bool" where |
23970 | 263 |
"CR1 \<tau> \<equiv> \<forall>t. (t\<in>RED \<tau> \<longrightarrow> SN t)" |
18106 | 264 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
265 |
definition "CR2" :: "ty \<Rightarrow> bool" where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
266 |
"CR2 \<tau> \<equiv> \<forall>t t'. (t\<in>RED \<tau> \<and> t \<longrightarrow>\<^sub>\<beta> t') \<longrightarrow> t'\<in>RED \<tau>" |
18106 | 267 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
268 |
definition "CR3_RED" :: "lam \<Rightarrow> ty \<Rightarrow> bool" where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
269 |
"CR3_RED t \<tau> \<equiv> \<forall>t'. t\<longrightarrow>\<^sub>\<beta> t' \<longrightarrow> t'\<in>RED \<tau>" |
18106 | 270 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
271 |
definition "CR3" :: "ty \<Rightarrow> bool" where |
18383 | 272 |
"CR3 \<tau> \<equiv> \<forall>t. (NEUT t \<and> CR3_RED t \<tau>) \<longrightarrow> t\<in>RED \<tau>" |
18106 | 273 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32960
diff
changeset
|
274 |
definition "CR4" :: "ty \<Rightarrow> bool" where |
18383 | 275 |
"CR4 \<tau> \<equiv> \<forall>t. (NEUT t \<and> NORMAL t) \<longrightarrow>t\<in>RED \<tau>" |
18106 | 276 |
|
23970 | 277 |
lemma CR3_implies_CR4: |
278 |
assumes a: "CR3 \<tau>" |
|
279 |
shows "CR4 \<tau>" |
|
280 |
using a by (auto simp add: CR3_def CR3_RED_def CR4_def NORMAL_def) |
|
18106 | 281 |
|
23970 | 282 |
(* sub_induction in the arrow-type case for the next proof *) |
283 |
lemma sub_induction: |
|
284 |
assumes a: "SN(u)" |
|
285 |
and b: "u\<in>RED \<tau>" |
|
286 |
and c1: "NEUT t" |
|
287 |
and c2: "CR2 \<tau>" |
|
288 |
and c3: "CR3 \<sigma>" |
|
289 |
and c4: "CR3_RED t (\<tau>\<rightarrow>\<sigma>)" |
|
290 |
shows "(App t u)\<in>RED \<sigma>" |
|
291 |
using a b |
|
292 |
proof (induct) |
|
293 |
fix u |
|
294 |
assume as: "u\<in>RED \<tau>" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
295 |
assume ih: " \<And>u'. \<lbrakk>u \<longrightarrow>\<^sub>\<beta> u'; u' \<in> RED \<tau>\<rbrakk> \<Longrightarrow> App t u' \<in> RED \<sigma>" |
23970 | 296 |
have "NEUT (App t u)" using c1 by (auto simp add: NEUT_def) |
297 |
moreover |
|
298 |
have "CR3_RED (App t u) \<sigma>" unfolding CR3_RED_def |
|
299 |
proof (intro strip) |
|
300 |
fix r |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
301 |
assume red: "App t u \<longrightarrow>\<^sub>\<beta> r" |
23970 | 302 |
moreover |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
303 |
{ assume "\<exists>t'. t \<longrightarrow>\<^sub>\<beta> t' \<and> r = App t' u" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
304 |
then obtain t' where a1: "t \<longrightarrow>\<^sub>\<beta> t'" and a2: "r = App t' u" by blast |
23970 | 305 |
have "t'\<in>RED (\<tau>\<rightarrow>\<sigma>)" using c4 a1 by (simp add: CR3_RED_def) |
306 |
then have "App t' u\<in>RED \<sigma>" using as by simp |
|
307 |
then have "r\<in>RED \<sigma>" using a2 by simp |
|
308 |
} |
|
309 |
moreover |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
310 |
{ assume "\<exists>u'. u \<longrightarrow>\<^sub>\<beta> u' \<and> r = App t u'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
311 |
then obtain u' where b1: "u \<longrightarrow>\<^sub>\<beta> u'" and b2: "r = App t u'" by blast |
23970 | 312 |
have "u'\<in>RED \<tau>" using as b1 c2 by (auto simp add: CR2_def) |
313 |
with ih have "App t u' \<in> RED \<sigma>" using b1 by simp |
|
314 |
then have "r\<in>RED \<sigma>" using b2 by simp |
|
315 |
} |
|
316 |
moreover |
|
317 |
{ assume "\<exists>x t'. t = Lam [x].t'" |
|
318 |
then obtain x t' where "t = Lam [x].t'" by blast |
|
319 |
then have "NEUT (Lam [x].t')" using c1 by simp |
|
320 |
then have "False" by (simp add: NEUT_def) |
|
321 |
then have "r\<in>RED \<sigma>" by simp |
|
322 |
} |
|
323 |
ultimately show "r \<in> RED \<sigma>" by (cases) (auto simp add: lam.inject) |
|
324 |
qed |
|
325 |
ultimately show "App t u \<in> RED \<sigma>" using c3 by (simp add: CR3_def) |
|
326 |
qed |
|
18106 | 327 |
|
63167 | 328 |
text \<open>properties of the candiadates\<close> |
18383 | 329 |
lemma RED_props: |
330 |
shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>" |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
331 |
proof (nominal_induct \<tau> rule: ty.strong_induct) |
18611 | 332 |
case (TVar a) |
333 |
{ case 1 show "CR1 (TVar a)" by (simp add: CR1_def) |
|
334 |
next |
|
23970 | 335 |
case 2 show "CR2 (TVar a)" by (auto intro: SN_preserved simp add: CR2_def) |
18611 | 336 |
next |
23970 | 337 |
case 3 show "CR3 (TVar a)" by (auto intro: SN_intro simp add: CR3_def CR3_RED_def) |
18611 | 338 |
} |
18599
e01112713fdc
changed PRO_RED proof to conform with the new induction rules
urbanc
parents:
18383
diff
changeset
|
339 |
next |
18611 | 340 |
case (TArr \<tau>1 \<tau>2) |
341 |
{ case 1 |
|
342 |
have ih_CR3_\<tau>1: "CR3 \<tau>1" by fact |
|
343 |
have ih_CR1_\<tau>2: "CR1 \<tau>2" by fact |
|
25867 | 344 |
have "\<And>t. t \<in> RED (\<tau>1 \<rightarrow> \<tau>2) \<Longrightarrow> SN t" |
345 |
proof - |
|
18611 | 346 |
fix t |
25867 | 347 |
assume "t \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" |
348 |
then have a: "\<forall>u. u \<in> RED \<tau>1 \<longrightarrow> App t u \<in> RED \<tau>2" by simp |
|
23970 | 349 |
from ih_CR3_\<tau>1 have "CR4 \<tau>1" by (simp add: CR3_implies_CR4) |
18611 | 350 |
moreover |
26932 | 351 |
fix a have "NEUT (Var a)" by (force simp add: NEUT_def) |
18611 | 352 |
moreover |
353 |
have "NORMAL (Var a)" by (rule NORMAL_Var) |
|
354 |
ultimately have "(Var a)\<in> RED \<tau>1" by (simp add: CR4_def) |
|
355 |
with a have "App t (Var a) \<in> RED \<tau>2" by simp |
|
356 |
hence "SN (App t (Var a))" using ih_CR1_\<tau>2 by (simp add: CR1_def) |
|
25867 | 357 |
thus "SN t" by (auto dest: SN_of_FST_of_App) |
18611 | 358 |
qed |
25867 | 359 |
then show "CR1 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR1_def by simp |
18611 | 360 |
next |
361 |
case 2 |
|
362 |
have ih_CR2_\<tau>2: "CR2 \<tau>2" by fact |
|
25867 | 363 |
then show "CR2 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR2_def by auto |
18611 | 364 |
next |
365 |
case 3 |
|
366 |
have ih_CR1_\<tau>1: "CR1 \<tau>1" by fact |
|
367 |
have ih_CR2_\<tau>1: "CR2 \<tau>1" by fact |
|
368 |
have ih_CR3_\<tau>2: "CR3 \<tau>2" by fact |
|
23970 | 369 |
show "CR3 (\<tau>1 \<rightarrow> \<tau>2)" unfolding CR3_def |
370 |
proof (simp, intro strip) |
|
18611 | 371 |
fix t u |
372 |
assume a1: "u \<in> RED \<tau>1" |
|
373 |
assume a2: "NEUT t \<and> CR3_RED t (\<tau>1 \<rightarrow> \<tau>2)" |
|
23970 | 374 |
have "SN(u)" using a1 ih_CR1_\<tau>1 by (simp add: CR1_def) |
375 |
then show "(App t u)\<in>RED \<tau>2" using ih_CR2_\<tau>1 ih_CR3_\<tau>2 a1 a2 by (blast intro: sub_induction) |
|
18611 | 376 |
qed |
377 |
} |
|
18383 | 378 |
qed |
23970 | 379 |
|
63167 | 380 |
text \<open> |
25867 | 381 |
the next lemma not as simple as on paper, probably because of |
382 |
the stronger double_SN induction |
|
63167 | 383 |
\<close> |
23970 | 384 |
lemma abs_RED: |
385 |
assumes asm: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>" |
|
386 |
shows "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)" |
|
18106 | 387 |
proof - |
23970 | 388 |
have b1: "SN t" |
389 |
proof - |
|
390 |
have "Var x\<in>RED \<tau>" |
|
391 |
proof - |
|
392 |
have "CR4 \<tau>" by (simp add: RED_props CR3_implies_CR4) |
|
393 |
moreover |
|
394 |
have "NEUT (Var x)" by (auto simp add: NEUT_def) |
|
395 |
moreover |
|
396 |
have "NORMAL (Var x)" by (auto elim: Beta.cases simp add: NORMAL_def) |
|
397 |
ultimately show "Var x\<in>RED \<tau>" by (simp add: CR4_def) |
|
398 |
qed |
|
399 |
then have "t[x::=Var x]\<in>RED \<sigma>" using asm by simp |
|
400 |
then have "t\<in>RED \<sigma>" by (simp add: id_subs) |
|
401 |
moreover |
|
402 |
have "CR1 \<sigma>" by (simp add: RED_props) |
|
403 |
ultimately show "SN t" by (simp add: CR1_def) |
|
404 |
qed |
|
405 |
show "Lam [x].t\<in>RED (\<tau>\<rightarrow>\<sigma>)" |
|
406 |
proof (simp, intro strip) |
|
407 |
fix u |
|
408 |
assume b2: "u\<in>RED \<tau>" |
|
409 |
then have b3: "SN u" using RED_props by (auto simp add: CR1_def) |
|
410 |
show "App (Lam [x].t) u \<in> RED \<sigma>" using b1 b3 b2 asm |
|
411 |
proof(induct t u rule: double_SN) |
|
412 |
fix t u |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
413 |
assume ih1: "\<And>t'. \<lbrakk>t \<longrightarrow>\<^sub>\<beta> t'; u\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t'[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t') u \<in> RED \<sigma>" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
414 |
assume ih2: "\<And>u'. \<lbrakk>u \<longrightarrow>\<^sub>\<beta> u'; u'\<in>RED \<tau>; \<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>\<rbrakk> \<Longrightarrow> App (Lam [x].t) u' \<in> RED \<sigma>" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
415 |
assume u: "u \<in> RED \<tau>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
416 |
assume t: "\<forall>s\<in>RED \<tau>. t[x::=s]\<in>RED \<sigma>" |
23970 | 417 |
have "CR3_RED (App (Lam [x].t) u) \<sigma>" unfolding CR3_RED_def |
418 |
proof(intro strip) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
419 |
fix r |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
420 |
assume red: "App (Lam [x].t) u \<longrightarrow>\<^sub>\<beta> r" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
421 |
moreover |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
422 |
{ assume "\<exists>t'. t \<longrightarrow>\<^sub>\<beta> t' \<and> r = App (Lam [x].t') u" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
423 |
then obtain t' where "t \<longrightarrow>\<^sub>\<beta> t'" and t': "r = App (Lam [x].t') u" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
424 |
by blast |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
425 |
then have "\<forall>s\<in>RED \<tau>. t'[x::=s] \<in> RED \<sigma>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
426 |
using CR2_def RED_props(2) t beta_subst by blast |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
427 |
then have "App (Lam [x].t') u\<in>RED \<sigma>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
428 |
using \<open>t \<longrightarrow>\<^sub>\<beta> t'\<close> u ih1 by blast |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
429 |
then have "r\<in>RED \<sigma>" using t' by simp |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
430 |
} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
431 |
moreover |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
35416
diff
changeset
|
432 |
{ assume "\<exists>u'. u \<longrightarrow>\<^sub>\<beta> u' \<and> r = App (Lam [x].t) u'" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
433 |
then obtain u' where "u \<longrightarrow>\<^sub>\<beta> u'" and u': "r = App (Lam [x].t) u'" by blast |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
434 |
have "CR2 \<tau>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
435 |
by (simp add: RED_props(2)) |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
436 |
then |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
437 |
have "App (Lam [x].t) u'\<in>RED \<sigma>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
438 |
using CR2_def ih2 \<open>u \<longrightarrow>\<^sub>\<beta> u'\<close> t u by blast |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
439 |
then have "r\<in>RED \<sigma>" using u' by simp |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
440 |
} |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
441 |
moreover |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
442 |
{ assume "r = t[x::=u]" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
443 |
then have "r\<in>RED \<sigma>" using u t by auto |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
444 |
} |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
445 |
ultimately show "r \<in> RED \<sigma>" |
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
446 |
by cases (auto simp: alpha subst_rename lam.inject dest: beta_abs) |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
447 |
(* one wants to use the strong elimination principle; for this one |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
448 |
has to know that x\<sharp>u *) |
18106 | 449 |
qed |
23970 | 450 |
moreover |
451 |
have "NEUT (App (Lam [x].t) u)" unfolding NEUT_def by (auto) |
|
452 |
ultimately show "App (Lam [x].t) u \<in> RED \<sigma>" using RED_props by (simp add: CR3_def) |
|
18106 | 453 |
qed |
454 |
qed |
|
23970 | 455 |
qed |
18106 | 456 |
|
22420 | 457 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
458 |
mapsto :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> bool" (\<open>_ maps _ to _\<close> [55,55,55] 55) |
22420 | 459 |
where |
25867 | 460 |
"\<theta> maps x to e \<equiv> (lookup \<theta> x) = e" |
22420 | 461 |
|
462 |
abbreviation |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80143
diff
changeset
|
463 |
closes :: "(name\<times>lam) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (\<open>_ closes _\<close> [55,55] 55) |
22420 | 464 |
where |
465 |
"\<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
|
466 |
|
18106 | 467 |
lemma all_RED: |
22420 | 468 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
469 |
and b: "\<theta> closes \<Gamma>" |
|
470 |
shows "\<theta><t> \<in> RED \<tau>" |
|
18345 | 471 |
using a b |
23142 | 472 |
proof(nominal_induct avoiding: \<theta> rule: typing.strong_induct) |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
63167
diff
changeset
|
473 |
case (t3 a \<Gamma> \<sigma> t \<tau> \<theta>) \<comment> \<open>lambda case\<close> |
23142 | 474 |
have ih: "\<And>\<theta>. \<theta> closes ((a,\<sigma>)#\<Gamma>) \<Longrightarrow> \<theta><t> \<in> RED \<tau>" by fact |
475 |
have \<theta>_cond: "\<theta> closes \<Gamma>" by fact |
|
23393 | 476 |
have fresh: "a\<sharp>\<Gamma>" "a\<sharp>\<theta>" by fact+ |
24899 | 477 |
from ih have "\<forall>s\<in>RED \<sigma>. ((a,s)#\<theta>)<t> \<in> RED \<tau>" using fresh \<theta>_cond fresh_context by simp |
478 |
then have "\<forall>s\<in>RED \<sigma>. \<theta><t>[a::=s] \<in> RED \<tau>" using fresh by (simp add: psubst_subst) |
|
23970 | 479 |
then have "Lam [a].(\<theta><t>) \<in> RED (\<sigma> \<rightarrow> \<tau>)" by (simp only: abs_RED) |
23142 | 480 |
then show "\<theta><(Lam [a].t)> \<in> RED (\<sigma> \<rightarrow> \<tau>)" using fresh by simp |
481 |
qed auto |
|
18345 | 482 |
|
63167 | 483 |
section \<open>identity substitution generated from a context \<Gamma>\<close> |
23142 | 484 |
fun |
18382 | 485 |
"id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list" |
23142 | 486 |
where |
18382 | 487 |
"id [] = []" |
23142 | 488 |
| "id ((x,\<tau>)#\<Gamma>) = (x,Var x)#(id \<Gamma>)" |
18382 | 489 |
|
23142 | 490 |
lemma id_maps: |
491 |
shows "(id \<Gamma>) maps a to (Var a)" |
|
492 |
by (induct \<Gamma>) (auto) |
|
18382 | 493 |
|
494 |
lemma id_fresh: |
|
495 |
fixes a::"name" |
|
496 |
assumes a: "a\<sharp>\<Gamma>" |
|
497 |
shows "a\<sharp>(id \<Gamma>)" |
|
498 |
using a |
|
23142 | 499 |
by (induct \<Gamma>) |
500 |
(auto simp add: fresh_list_nil fresh_list_cons) |
|
18382 | 501 |
|
502 |
lemma id_apply: |
|
22420 | 503 |
shows "(id \<Gamma>)<t> = t" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26932
diff
changeset
|
504 |
by (nominal_induct t avoiding: \<Gamma> rule: lam.strong_induct) |
23970 | 505 |
(auto simp add: id_maps id_fresh) |
18382 | 506 |
|
23142 | 507 |
lemma id_closes: |
22420 | 508 |
shows "(id \<Gamma>) closes \<Gamma>" |
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
509 |
by (metis CR3_implies_CR4 CR4_def NEUT_def NORMAL_Var RED_props(3) id_maps) |
18382 | 510 |
|
80143
378593bf5109
Tidying up another of the nominal examples
paulson <lp15@cam.ac.uk>
parents:
67443
diff
changeset
|
511 |
lemma typing_implies_RED: |
23142 | 512 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
18383 | 513 |
shows "t \<in> RED \<tau>" |
514 |
proof - |
|
22420 | 515 |
have "(id \<Gamma>)<t>\<in>RED \<tau>" |
18383 | 516 |
proof - |
23142 | 517 |
have "(id \<Gamma>) closes \<Gamma>" by (rule id_closes) |
18383 | 518 |
with a show ?thesis by (rule all_RED) |
519 |
qed |
|
520 |
thus"t \<in> RED \<tau>" by (simp add: id_apply) |
|
521 |
qed |
|
522 |
||
523 |
lemma typing_implies_SN: |
|
23142 | 524 |
assumes a: "\<Gamma> \<turnstile> t : \<tau>" |
18383 | 525 |
shows "SN(t)" |
526 |
proof - |
|
527 |
from a have "t \<in> RED \<tau>" by (rule typing_implies_RED) |
|
528 |
moreover |
|
529 |
have "CR1 \<tau>" by (rule RED_props) |
|
530 |
ultimately show "SN(t)" by (simp add: CR1_def) |
|
531 |
qed |
|
18382 | 532 |
|
62390 | 533 |
end |