author | paulson <lp15@cam.ac.uk> |
Mon, 22 Apr 2024 10:43:57 +0100 | |
changeset 80141 | 022a9c26b14f |
parent 71989 | bad75618fb82 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
27623 | 1 |
(* Title: HOL/Nominal/Examples/Standardization.thy |
2 |
Author: Stefan Berghofer and Tobias Nipkow |
|
3 |
Copyright 2005, 2008 TU Muenchen |
|
4 |
*) |
|
5 |
||
63167 | 6 |
section \<open>Standardization\<close> |
27623 | 7 |
|
8 |
theory Standardization |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63882
diff
changeset
|
9 |
imports "HOL-Nominal.Nominal" |
27623 | 10 |
begin |
11 |
||
63167 | 12 |
text \<open> |
27623 | 13 |
The proof of the standardization theorem, as well as most of the theorems about |
63167 | 14 |
lambda calculus in the following sections, are taken from \<open>HOL/Lambda\<close>. |
15 |
\<close> |
|
27623 | 16 |
|
63167 | 17 |
subsection \<open>Lambda terms\<close> |
27623 | 18 |
|
19 |
atom_decl name |
|
20 |
||
21 |
nominal_datatype lam = |
|
22 |
Var "name" |
|
23 |
| App "lam" "lam" (infixl "\<degree>" 200) |
|
24 |
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [0, 10] 10) |
|
25 |
||
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
26 |
instantiation lam :: size |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
27 |
begin |
27623 | 28 |
|
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
29 |
nominal_primrec size_lam |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
30 |
where |
27623 | 31 |
"size (Var n) = 0" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
32 |
| "size (t \<degree> u) = size t + size u + 1" |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
33 |
| "size (Lam [x].t) = size t + 1" |
80141 | 34 |
by (finite_guess | simp add: fresh_nat | fresh_guess)+ |
27623 | 35 |
|
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
36 |
instance .. |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
37 |
|
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
38 |
end |
27623 | 39 |
|
40 |
nominal_primrec |
|
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
41 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [300, 0, 0] 300) |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
42 |
where |
27623 | 43 |
subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
47397
diff
changeset
|
44 |
| subst_App: "(t\<^sub>1 \<degree> t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \<degree> t\<^sub>2[y::=s]" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
27623
diff
changeset
|
45 |
| subst_Lam: "x \<sharp> (y, s) \<Longrightarrow> (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))" |
80141 | 46 |
by (finite_guess | simp add: abs_fresh | fresh_guess)+ |
27623 | 47 |
|
48 |
lemma subst_eqvt [eqvt]: |
|
49 |
"(pi::name prm) \<bullet> (t[x::=u]) = (pi \<bullet> t)[(pi \<bullet> x)::=(pi \<bullet> u)]" |
|
50 |
by (nominal_induct t avoiding: x u rule: lam.strong_induct) |
|
80141 | 51 |
(perm_simp add: fresh_bij)+ |
27623 | 52 |
|
53 |
lemma subst_rename: |
|
54 |
"y \<sharp> t \<Longrightarrow> ([(y, x)] \<bullet> t)[y::=u] = t[x::=u]" |
|
55 |
by (nominal_induct t avoiding: x y u rule: lam.strong_induct) |
|
80141 | 56 |
(simp_all add: fresh_atm calc_atm abs_fresh) |
27623 | 57 |
|
58 |
lemma fresh_subst: |
|
59 |
"(x::name) \<sharp> t \<Longrightarrow> x \<sharp> u \<Longrightarrow> x \<sharp> t[y::=u]" |
|
60 |
by (nominal_induct t avoiding: x y u rule: lam.strong_induct) |
|
80141 | 61 |
(auto simp add: abs_fresh fresh_atm) |
27623 | 62 |
|
63 |
lemma fresh_subst': |
|
64 |
"(x::name) \<sharp> u \<Longrightarrow> x \<sharp> t[x::=u]" |
|
65 |
by (nominal_induct t avoiding: x u rule: lam.strong_induct) |
|
80141 | 66 |
(auto simp add: abs_fresh fresh_atm) |
27623 | 67 |
|
68 |
lemma subst_forget: "(x::name) \<sharp> t \<Longrightarrow> t[x::=u] = t" |
|
69 |
by (nominal_induct t avoiding: x u rule: lam.strong_induct) |
|
80141 | 70 |
(auto simp add: abs_fresh fresh_atm) |
27623 | 71 |
|
72 |
lemma subst_subst: |
|
73 |
"x \<noteq> y \<Longrightarrow> x \<sharp> v \<Longrightarrow> t[y::=v][x::=u[y::=v]] = t[x::=u][y::=v]" |
|
74 |
by (nominal_induct t avoiding: x y u v rule: lam.strong_induct) |
|
80141 | 75 |
(auto simp add: fresh_subst subst_forget) |
27623 | 76 |
|
77 |
declare subst_Var [simp del] |
|
78 |
||
79 |
lemma subst_eq [simp]: "(Var x)[x::=u] = u" |
|
80 |
by (simp add: subst_Var) |
|
81 |
||
82 |
lemma subst_neq [simp]: "x \<noteq> y \<Longrightarrow> (Var x)[y::=u] = Var x" |
|
83 |
by (simp add: subst_Var) |
|
84 |
||
85 |
inductive beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50) |
|
86 |
where |
|
87 |
beta: "x \<sharp> t \<Longrightarrow> (Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]" |
|
88 |
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u" |
|
89 |
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t" |
|
90 |
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta> (Lam [x].t)" |
|
91 |
||
92 |
equivariance beta |
|
93 |
nominal_inductive beta |
|
94 |
by (simp_all add: abs_fresh fresh_subst') |
|
95 |
||
96 |
lemma better_beta [simp, intro!]: "(Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]" |
|
97 |
proof - |
|
98 |
obtain y::name where y: "y \<sharp> (x, s, t)" |
|
99 |
by (rule exists_fresh) (rule fin_supp) |
|
100 |
then have "y \<sharp> t" by simp |
|
101 |
then have "(Lam [y]. [(y, x)] \<bullet> s) \<degree> t \<rightarrow>\<^sub>\<beta> ([(y, x)] \<bullet> s)[y::=t]" |
|
102 |
by (rule beta) |
|
103 |
moreover from y have "(Lam [x].s) = (Lam [y]. [(y, x)] \<bullet> s)" |
|
104 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
105 |
ultimately show ?thesis using y by (simp add: subst_rename) |
|
106 |
qed |
|
107 |
||
108 |
abbreviation |
|
109 |
beta_reds :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50) where |
|
110 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t \<equiv> beta\<^sup>*\<^sup>* s t" |
|
111 |
||
112 |
||
63167 | 113 |
subsection \<open>Application of a term to a list of terms\<close> |
27623 | 114 |
|
115 |
abbreviation |
|
116 |
list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam" (infixl "\<degree>\<degree>" 150) where |
|
67399 | 117 |
"t \<degree>\<degree> ts \<equiv> foldl (\<degree>) t ts" |
27623 | 118 |
|
119 |
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)" |
|
120 |
by (induct ts rule: rev_induct) (auto simp add: lam.inject) |
|
121 |
||
122 |
lemma Var_eq_apps_conv [iff]: "(Var m = s \<degree>\<degree> ss) = (Var m = s \<and> ss = [])" |
|
123 |
by (induct ss arbitrary: s) auto |
|
124 |
||
125 |
lemma Var_apps_eq_Var_apps_conv [iff]: |
|
126 |
"(Var m \<degree>\<degree> rs = Var n \<degree>\<degree> ss) = (m = n \<and> rs = ss)" |
|
80141 | 127 |
proof (induct rs arbitrary: ss rule: rev_induct) |
128 |
case Nil then show ?case by (auto simp add: lam.inject) |
|
129 |
next |
|
130 |
case (snoc x xs) then show ?case |
|
131 |
by (induct ss rule: rev_induct) (auto simp add: lam.inject) |
|
132 |
qed |
|
27623 | 133 |
|
134 |
lemma App_eq_foldl_conv: |
|
135 |
"(r \<degree> s = t \<degree>\<degree> ts) = |
|
136 |
(if ts = [] then r \<degree> s = t |
|
137 |
else (\<exists>ss. ts = ss @ [s] \<and> r = t \<degree>\<degree> ss))" |
|
80141 | 138 |
by (rule rev_exhaust [of ts]) (auto simp: lam.inject) |
27623 | 139 |
|
140 |
lemma Abs_eq_apps_conv [iff]: |
|
141 |
"((Lam [x].r) = s \<degree>\<degree> ss) = ((Lam [x].r) = s \<and> ss = [])" |
|
142 |
by (induct ss rule: rev_induct) auto |
|
143 |
||
144 |
lemma apps_eq_Abs_conv [iff]: "(s \<degree>\<degree> ss = (Lam [x].r)) = (s = (Lam [x].r) \<and> ss = [])" |
|
145 |
by (induct ss rule: rev_induct) auto |
|
146 |
||
147 |
lemma Abs_App_neq_Var_apps [iff]: |
|
148 |
"(Lam [x].s) \<degree> t \<noteq> Var n \<degree>\<degree> ss" |
|
149 |
by (induct ss arbitrary: s t rule: rev_induct) (auto simp add: lam.inject) |
|
150 |
||
151 |
lemma Var_apps_neq_Abs_apps [iff]: |
|
152 |
"Var n \<degree>\<degree> ts \<noteq> (Lam [x].r) \<degree>\<degree> ss" |
|
80141 | 153 |
proof (induct ss arbitrary: ts rule: rev_induct) |
154 |
case Nil then show ?case by simp |
|
155 |
next |
|
156 |
case (snoc x xs) then show ?case |
|
157 |
by (induct ts rule: rev_induct) (auto simp add: lam.inject) |
|
158 |
qed |
|
27623 | 159 |
|
160 |
lemma ex_head_tail: |
|
161 |
"\<exists>ts h. t = h \<degree>\<degree> ts \<and> ((\<exists>n. h = Var n) \<or> (\<exists>x u. h = (Lam [x].u)))" |
|
80141 | 162 |
proof (induct t rule: lam.induct) |
163 |
case (App lam1 lam2) |
|
164 |
then show ?case |
|
165 |
by (metis foldl_Cons foldl_Nil foldl_append) |
|
166 |
qed auto |
|
27623 | 167 |
|
168 |
lemma size_apps [simp]: |
|
67399 | 169 |
"size (r \<degree>\<degree> rs) = size r + foldl (+) 0 (map size rs) + length rs" |
27623 | 170 |
by (induct rs rule: rev_induct) auto |
171 |
||
172 |
lemma lem0: "(0::nat) < k \<Longrightarrow> m \<le> n \<Longrightarrow> m < n + k" |
|
173 |
by simp |
|
174 |
||
175 |
lemma subst_map [simp]: |
|
176 |
"(t \<degree>\<degree> ts)[x::=u] = t[x::=u] \<degree>\<degree> map (\<lambda>t. t[x::=u]) ts" |
|
177 |
by (induct ts arbitrary: t) simp_all |
|
178 |
||
179 |
lemma app_last: "(t \<degree>\<degree> ts) \<degree> u = t \<degree>\<degree> (ts @ [u])" |
|
180 |
by simp |
|
181 |
||
182 |
lemma perm_apps [eqvt]: |
|
183 |
"(pi::name prm) \<bullet> (t \<degree>\<degree> ts) = ((pi \<bullet> t) \<degree>\<degree> (pi \<bullet> ts))" |
|
184 |
by (induct ts rule: rev_induct) (auto simp add: append_eqvt) |
|
185 |
||
186 |
lemma fresh_apps [simp]: "(x::name) \<sharp> (t \<degree>\<degree> ts) = (x \<sharp> t \<and> x \<sharp> ts)" |
|
187 |
by (induct ts rule: rev_induct) |
|
188 |
(auto simp add: fresh_list_append fresh_list_nil fresh_list_cons) |
|
189 |
||
63167 | 190 |
text \<open>A customized induction schema for \<open>\<degree>\<degree>\<close>.\<close> |
27623 | 191 |
|
80141 | 192 |
lemma Apps_lam_induct_aux: |
27623 | 193 |
assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)" |
194 |
and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)" |
|
195 |
shows "size t = n \<Longrightarrow> P z t" |
|
80141 | 196 |
proof (induct n arbitrary: t z rule: less_induct) |
197 |
case (less n) |
|
198 |
obtain ts h where t: "t = h \<degree>\<degree> ts" and D: "(\<exists>a. h = Var a) \<or> (\<exists>x u. h = (Lam [x].u))" |
|
199 |
using ex_head_tail [of t] by metis |
|
200 |
show ?case |
|
201 |
using D |
|
202 |
proof (elim exE disjE) |
|
203 |
fix a :: name |
|
204 |
assume h: "h = Var a" |
|
205 |
have "P z t" if "t \<in> set ts" for z t |
|
206 |
proof - |
|
207 |
have "size t < length ts + fold (+) (map size ts) 0" |
|
208 |
using that |
|
209 |
by (fastforce simp add: sum_list_map_remove1 fold_plus_sum_list_rev) |
|
210 |
then have "size t < size (Var a \<degree>\<degree> ts)" |
|
211 |
by simp (simp add: add.commute foldl_conv_fold) |
|
212 |
then show ?thesis |
|
213 |
using h less.hyps less.prems t by blast |
|
214 |
qed |
|
215 |
then show "P z t" |
|
216 |
unfolding t h by (blast intro: assms) |
|
217 |
next |
|
218 |
fix x u |
|
219 |
assume h: "h = (Lam [x].u)" |
|
220 |
obtain y::name where y: "y \<sharp> (x, u, z)" |
|
221 |
by (metis exists_fresh' fin_supp) |
|
222 |
then have eq: "(Lam [x].u) = (Lam [y].([(y, x)] \<bullet> u))" |
|
223 |
by (metis alpha' fresh_prod lam.inject(3) perm_fresh_fresh) |
|
224 |
show "P z t" |
|
225 |
unfolding t h eq |
|
226 |
proof (intro assms strip) |
|
227 |
show "y \<sharp> z" |
|
228 |
by (simp add: y) |
|
229 |
have "size ([(y, x)] \<bullet> u) < size ((Lam [x].u) \<degree>\<degree> ts)" |
|
230 |
by (simp add: eq) |
|
231 |
then show "P z ([(y, x)] \<bullet> u)" for z |
|
232 |
using h less.hyps less.prems t by blast |
|
233 |
show "P z t" if "t\<in>set ts" for z t |
|
234 |
proof - |
|
235 |
have 2: "size t < size ((Lam [x].u) \<degree>\<degree> ts)" |
|
236 |
using that |
|
237 |
apply (simp add: eq) |
|
238 |
apply (simp only: foldl_conv_fold add.commute fold_plus_sum_list_rev) |
|
239 |
apply (fastforce simp add: sum_list_map_remove1) |
|
240 |
done |
|
241 |
then show ?thesis |
|
242 |
using h less.hyps less.prems t by blast |
|
243 |
qed |
|
244 |
qed |
|
245 |
qed |
|
246 |
qed |
|
27623 | 247 |
|
248 |
theorem Apps_lam_induct: |
|
249 |
assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)" |
|
250 |
and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)" |
|
251 |
shows "P z t" |
|
80141 | 252 |
using Apps_lam_induct_aux [of P] assms by blast |
27623 | 253 |
|
254 |
||
63167 | 255 |
subsection \<open>Congruence rules\<close> |
27623 | 256 |
|
257 |
lemma apps_preserves_beta [simp]: |
|
258 |
"r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> r \<degree>\<degree> ss \<rightarrow>\<^sub>\<beta> s \<degree>\<degree> ss" |
|
259 |
by (induct ss rule: rev_induct) auto |
|
260 |
||
261 |
lemma rtrancl_beta_Abs [intro!]: |
|
262 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta>\<^sup>* (Lam [x].s')" |
|
263 |
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ |
|
264 |
||
265 |
lemma rtrancl_beta_AppL: |
|
266 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t" |
|
267 |
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ |
|
268 |
||
269 |
lemma rtrancl_beta_AppR: |
|
270 |
"t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s \<degree> t'" |
|
271 |
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ |
|
272 |
||
273 |
lemma rtrancl_beta_App [intro]: |
|
274 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'" |
|
275 |
by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtranclp_trans) |
|
276 |
||
277 |
||
63167 | 278 |
subsection \<open>Lifting an order to lists of elements\<close> |
27623 | 279 |
|
280 |
definition |
|
281 |
step1 :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where |
|
80141 | 282 |
"step1 r \<equiv> |
27623 | 283 |
(\<lambda>ys xs. \<exists>us z z' vs. xs = us @ z # vs \<and> r z' z \<and> ys = |
284 |
us @ z' # vs)" |
|
285 |
||
286 |
lemma not_Nil_step1 [iff]: "\<not> step1 r [] xs" |
|
80141 | 287 |
by (simp add: step1_def) |
27623 | 288 |
|
289 |
lemma not_step1_Nil [iff]: "\<not> step1 r xs []" |
|
80141 | 290 |
by (simp add: step1_def) |
27623 | 291 |
|
292 |
lemma Cons_step1_Cons [iff]: |
|
80141 | 293 |
"step1 r (y # ys) (x # xs) \<longleftrightarrow> r y x \<and> xs = ys \<or> x = y \<and> step1 r ys xs" |
294 |
apply (rule ) |
|
295 |
apply (smt (verit, ccfv_SIG) append_eq_Cons_conv list.inject step1_def) |
|
296 |
by (metis append_Cons append_Nil step1_def) |
|
27623 | 297 |
|
298 |
lemma Cons_step1E [elim!]: |
|
299 |
assumes "step1 r ys (x # xs)" |
|
300 |
and "\<And>y. ys = y # xs \<Longrightarrow> r y x \<Longrightarrow> R" |
|
301 |
and "\<And>zs. ys = x # zs \<Longrightarrow> step1 r zs xs \<Longrightarrow> R" |
|
302 |
shows R |
|
80141 | 303 |
by (metis Cons_step1_Cons assms list.exhaust not_Nil_step1) |
304 |
||
305 |
lemma append_step1I: |
|
306 |
"step1 r ys xs \<and> vs = us \<or> ys = xs \<and> step1 r vs us |
|
307 |
\<Longrightarrow> step1 r (ys @ vs) (xs @ us)" |
|
308 |
by (smt (verit) append_Cons append_assoc step1_def) |
|
27623 | 309 |
|
310 |
lemma Snoc_step1_SnocD: |
|
80141 | 311 |
assumes "step1 r (ys @ [y]) (xs @ [x])" |
312 |
shows "(step1 r ys xs \<and> y = x \<or> ys = xs \<and> r y x)" |
|
313 |
using assms |
|
314 |
apply (clarsimp simp: step1_def) |
|
315 |
by (metis butlast.simps(2) butlast_append butlast_snoc last.simps last_appendR list.simps(3)) |
|
27623 | 316 |
|
317 |
||
63167 | 318 |
subsection \<open>Lifting beta-reduction to lists\<close> |
27623 | 319 |
|
320 |
abbreviation |
|
321 |
list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>\<beta>]\<^sub>1" 50) where |
|
322 |
"rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<equiv> step1 beta rs ss" |
|
323 |
||
324 |
lemma head_Var_reduction: |
|
325 |
"Var n \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> v \<Longrightarrow> \<exists>ss. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<and> v = Var n \<degree>\<degree> ss" |
|
80141 | 326 |
proof (induct u \<equiv> "Var n \<degree>\<degree> rs" v arbitrary: rs set: beta) |
327 |
case (appL s t u) |
|
328 |
then show ?case |
|
329 |
by (smt (verit, best) App_eq_foldl_conv app_last append_step1I lam.distinct(1)) |
|
330 |
next |
|
331 |
case (appR s t u) |
|
332 |
then show ?case |
|
333 |
by (smt (verit, ccfv_SIG) App_eq_foldl_conv Cons_step1_Cons app_last append_step1I lam.distinct(1)) |
|
334 |
qed auto |
|
27623 | 335 |
|
336 |
lemma apps_betasE [case_names appL appR beta, consumes 1]: |
|
337 |
assumes major: "r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> s" |
|
338 |
and cases: "\<And>r'. r \<rightarrow>\<^sub>\<beta> r' \<Longrightarrow> s = r' \<degree>\<degree> rs \<Longrightarrow> R" |
|
339 |
"\<And>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<Longrightarrow> s = r \<degree>\<degree> rs' \<Longrightarrow> R" |
|
340 |
"\<And>t u us. (x \<sharp> r \<Longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us) \<Longrightarrow> R" |
|
341 |
shows R |
|
342 |
proof - |
|
80141 | 343 |
note [[simproc del: defined_all]] |
27623 | 344 |
from major have |
345 |
"(\<exists>r'. r \<rightarrow>\<^sub>\<beta> r' \<and> s = r' \<degree>\<degree> rs) \<or> |
|
346 |
(\<exists>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<and> s = r \<degree>\<degree> rs') \<or> |
|
347 |
(\<exists>t u us. x \<sharp> r \<longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us)" |
|
80141 | 348 |
proof (nominal_induct u \<equiv> "r \<degree>\<degree> rs" s avoiding: x r rs rule: beta.strong_induct) |
349 |
case (beta y t s) |
|
350 |
then show ?case |
|
351 |
apply (simp add: App_eq_foldl_conv split: if_split_asm) |
|
27623 | 352 |
apply blast |
80141 | 353 |
by (metis (no_types, lifting) abs_fresh(1) alpha' lam.fresh(3) lam.inject(3) subst_rename) |
354 |
next |
|
355 |
case (appL s t u) |
|
356 |
then show ?case |
|
357 |
apply (simp add: App_eq_foldl_conv split: if_split_asm) |
|
27623 | 358 |
apply blast |
80141 | 359 |
by (smt (verit) append_Cons append_step1I snoc_eq_iff_butlast) |
360 |
next |
|
361 |
case (appR s t u) |
|
362 |
then show ?case |
|
363 |
apply (simp add: App_eq_foldl_conv step1_def split: if_split_asm) |
|
364 |
apply force |
|
365 |
by (metis snoc_eq_iff_butlast) |
|
366 |
next |
|
367 |
case (abs s t x) |
|
368 |
then show ?case |
|
369 |
by blast |
|
370 |
qed |
|
27623 | 371 |
with cases show ?thesis by blast |
372 |
qed |
|
373 |
||
374 |
lemma apps_preserves_betas [simp]: |
|
375 |
"rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> r \<degree>\<degree> ss" |
|
80141 | 376 |
proof (induct rs arbitrary: ss rule: rev_induct) |
377 |
case Nil |
|
378 |
then show ?case by simp |
|
379 |
next |
|
380 |
case (snoc x ts) |
|
381 |
then show ?case |
|
382 |
apply (simp add: step1_def) |
|
383 |
by (smt (verit) appR app_last apps_preserves_beta foldl_Cons foldl_append) |
|
384 |
qed |
|
27623 | 385 |
|
386 |
||
63167 | 387 |
subsection \<open>Standard reduction relation\<close> |
27623 | 388 |
|
63167 | 389 |
text \<open> |
27623 | 390 |
Based on lecture notes by Ralph Matthes, |
391 |
original proof idea due to Ralph Loader. |
|
63167 | 392 |
\<close> |
27623 | 393 |
|
394 |
declare listrel_mono [mono_set] |
|
395 |
||
396 |
lemma listrelp_eqvt [eqvt]: |
|
46317 | 397 |
fixes f :: "'a::pt_name \<Rightarrow> 'b::pt_name \<Rightarrow> bool" |
27623 | 398 |
assumes xy: "listrelp f (x::'a::pt_name list) y" |
399 |
shows "listrelp ((pi::name prm) \<bullet> f) (pi \<bullet> x) (pi \<bullet> y)" using xy |
|
45966 | 400 |
by induct (simp_all add: listrelp.intros perm_app [symmetric]) |
27623 | 401 |
|
402 |
inductive |
|
403 |
sred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>s" 50) |
|
404 |
and sredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>s]" 50) |
|
405 |
where |
|
67399 | 406 |
"s [\<rightarrow>\<^sub>s] t \<equiv> listrelp (\<rightarrow>\<^sub>s) s t" |
27623 | 407 |
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'" |
408 |
| Abs: "x \<sharp> (ss, ss') \<Longrightarrow> r \<rightarrow>\<^sub>s r' \<Longrightarrow> ss [\<rightarrow>\<^sub>s] ss' \<Longrightarrow> (Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'" |
|
409 |
| Beta: "x \<sharp> (s, ss, t) \<Longrightarrow> r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
|
410 |
||
411 |
equivariance sred |
|
412 |
nominal_inductive sred |
|
413 |
by (simp add: abs_fresh)+ |
|
414 |
||
415 |
lemma better_sred_Abs: |
|
416 |
assumes H1: "r \<rightarrow>\<^sub>s r'" |
|
417 |
and H2: "ss [\<rightarrow>\<^sub>s] ss'" |
|
418 |
shows "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'" |
|
419 |
proof - |
|
420 |
obtain y::name where y: "y \<sharp> (x, r, r', ss, ss')" |
|
421 |
by (rule exists_fresh) (rule fin_supp) |
|
422 |
then have "y \<sharp> (ss, ss')" by simp |
|
423 |
moreover from H1 have "[(y, x)] \<bullet> (r \<rightarrow>\<^sub>s r')" by (rule perm_boolI) |
|
424 |
then have "([(y, x)] \<bullet> r) \<rightarrow>\<^sub>s ([(y, x)] \<bullet> r')" by (simp add: eqvts) |
|
425 |
ultimately have "(Lam [y]. [(y, x)] \<bullet> r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [y]. [(y, x)] \<bullet> r') \<degree>\<degree> ss'" using H2 |
|
426 |
by (rule sred.Abs) |
|
427 |
moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)" |
|
428 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
429 |
moreover from y have "(Lam [x].r') = (Lam [y]. [(y, x)] \<bullet> r')" |
|
430 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
431 |
ultimately show ?thesis by simp |
|
432 |
qed |
|
433 |
||
434 |
lemma better_sred_Beta: |
|
435 |
assumes H: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
|
436 |
shows "(Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
|
437 |
proof - |
|
438 |
obtain y::name where y: "y \<sharp> (x, r, s, ss, t)" |
|
439 |
by (rule exists_fresh) (rule fin_supp) |
|
440 |
then have "y \<sharp> (s, ss, t)" by simp |
|
441 |
moreover from y H have "([(y, x)] \<bullet> r)[y::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
|
442 |
by (simp add: subst_rename) |
|
443 |
ultimately have "(Lam [y].[(y, x)] \<bullet> r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t" |
|
444 |
by (rule sred.Beta) |
|
445 |
moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)" |
|
446 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
447 |
ultimately show ?thesis by simp |
|
448 |
qed |
|
449 |
||
450 |
lemmas better_sred_intros = sred.Var better_sred_Abs better_sred_Beta |
|
451 |
||
452 |
lemma refl_listrelp: "\<forall>x\<in>set xs. R x x \<Longrightarrow> listrelp R xs xs" |
|
453 |
by (induct xs) (auto intro: listrelp.intros) |
|
454 |
||
455 |
lemma refl_sred: "t \<rightarrow>\<^sub>s t" |
|
456 |
by (nominal_induct t rule: Apps_lam_induct) (auto intro: refl_listrelp better_sred_intros) |
|
457 |
||
458 |
lemma listrelp_conj1: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp R x y" |
|
459 |
by (erule listrelp.induct) (auto intro: listrelp.intros) |
|
460 |
||
461 |
lemma listrelp_conj2: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp S x y" |
|
462 |
by (erule listrelp.induct) (auto intro: listrelp.intros) |
|
463 |
||
464 |
lemma listrelp_app: |
|
465 |
assumes xsys: "listrelp R xs ys" |
|
466 |
shows "listrelp R xs' ys' \<Longrightarrow> listrelp R (xs @ xs') (ys @ ys')" using xsys |
|
467 |
by (induct arbitrary: xs' ys') (auto intro: listrelp.intros) |
|
468 |
||
469 |
lemma lemma1: |
|
470 |
assumes r: "r \<rightarrow>\<^sub>s r'" and s: "s \<rightarrow>\<^sub>s s'" |
|
471 |
shows "r \<degree> s \<rightarrow>\<^sub>s r' \<degree> s'" using r |
|
472 |
proof induct |
|
473 |
case (Var rs rs' x) |
|
474 |
then have "rs [\<rightarrow>\<^sub>s] rs'" by (rule listrelp_conj1) |
|
475 |
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros) |
|
476 |
ultimately have "rs @ [s] [\<rightarrow>\<^sub>s] rs' @ [s']" by (rule listrelp_app) |
|
477 |
hence "Var x \<degree>\<degree> (rs @ [s]) \<rightarrow>\<^sub>s Var x \<degree>\<degree> (rs' @ [s'])" by (rule sred.Var) |
|
478 |
thus ?case by (simp only: app_last) |
|
479 |
next |
|
480 |
case (Abs x ss ss' r r') |
|
481 |
from Abs(4) have "ss [\<rightarrow>\<^sub>s] ss'" by (rule listrelp_conj1) |
|
482 |
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros) |
|
483 |
ultimately have "ss @ [s] [\<rightarrow>\<^sub>s] ss' @ [s']" by (rule listrelp_app) |
|
63167 | 484 |
with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> (ss' @ [s'])" |
27623 | 485 |
by (rule better_sred_Abs) |
486 |
thus ?case by (simp only: app_last) |
|
487 |
next |
|
488 |
case (Beta x u ss t r) |
|
489 |
hence "r[x::=u] \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (simp only: app_last) |
|
490 |
hence "(Lam [x].r) \<degree> u \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (rule better_sred_Beta) |
|
491 |
thus ?case by (simp only: app_last) |
|
492 |
qed |
|
493 |
||
494 |
lemma lemma1': |
|
495 |
assumes ts: "ts [\<rightarrow>\<^sub>s] ts'" |
|
496 |
shows "r \<rightarrow>\<^sub>s r' \<Longrightarrow> r \<degree>\<degree> ts \<rightarrow>\<^sub>s r' \<degree>\<degree> ts'" using ts |
|
497 |
by (induct arbitrary: r r') (auto intro: lemma1) |
|
498 |
||
499 |
lemma listrelp_betas: |
|
67399 | 500 |
assumes ts: "listrelp (\<rightarrow>\<^sub>\<beta>\<^sup>*) ts ts'" |
27623 | 501 |
shows "\<And>t t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<degree>\<degree> ts'" using ts |
502 |
by induct auto |
|
503 |
||
504 |
lemma lemma2: |
|
505 |
assumes t: "t \<rightarrow>\<^sub>s u" |
|
506 |
shows "t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using t |
|
507 |
by induct (auto dest: listrelp_conj2 |
|
508 |
intro: listrelp_betas apps_preserves_beta converse_rtranclp_into_rtranclp) |
|
509 |
||
510 |
lemma lemma3: |
|
511 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
|
512 |
shows "s \<rightarrow>\<^sub>s s' \<Longrightarrow> r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" using r |
|
513 |
proof (nominal_induct avoiding: x s s' rule: sred.strong_induct) |
|
514 |
case (Var rs rs' y) |
|
515 |
hence "map (\<lambda>t. t[x::=s]) rs [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) rs'" |
|
516 |
by induct (auto intro: listrelp.intros Var) |
|
517 |
moreover have "Var y[x::=s] \<rightarrow>\<^sub>s Var y[x::=s']" |
|
518 |
by (cases "y = x") (auto simp add: Var intro: refl_sred) |
|
519 |
ultimately show ?case by simp (rule lemma1') |
|
520 |
next |
|
521 |
case (Abs y ss ss' r r') |
|
522 |
then have "r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" by fast |
|
63167 | 523 |
moreover from Abs(8) \<open>s \<rightarrow>\<^sub>s s'\<close> have "map (\<lambda>t. t[x::=s]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) ss'" |
27623 | 524 |
by induct (auto intro: listrelp.intros Abs) |
63167 | 525 |
ultimately show ?case using Abs(6) \<open>y \<sharp> x\<close> \<open>y \<sharp> s\<close> \<open>y \<sharp> s'\<close> |
27623 | 526 |
by simp (rule better_sred_Abs) |
527 |
next |
|
528 |
case (Beta y u ss t r) |
|
529 |
thus ?case by (auto simp add: subst_subst fresh_atm intro: better_sred_Beta) |
|
530 |
qed |
|
531 |
||
532 |
lemma lemma4_aux: |
|
533 |
assumes rs: "listrelp (\<lambda>t u. t \<rightarrow>\<^sub>s u \<and> (\<forall>r. u \<rightarrow>\<^sub>\<beta> r \<longrightarrow> t \<rightarrow>\<^sub>s r)) rs rs'" |
|
534 |
shows "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> rs [\<rightarrow>\<^sub>s] ss" using rs |
|
535 |
proof (induct arbitrary: ss) |
|
536 |
case Nil |
|
537 |
thus ?case by cases (auto intro: listrelp.Nil) |
|
538 |
next |
|
539 |
case (Cons x y xs ys) |
|
540 |
note Cons' = Cons |
|
541 |
show ?case |
|
542 |
proof (cases ss) |
|
543 |
case Nil with Cons show ?thesis by simp |
|
544 |
next |
|
545 |
case (Cons y' ys') |
|
546 |
hence ss: "ss = y' # ys'" by simp |
|
547 |
from Cons Cons' have "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys \<or> y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'" by simp |
|
548 |
hence "x # xs [\<rightarrow>\<^sub>s] y' # ys'" |
|
549 |
proof |
|
550 |
assume H: "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys" |
|
551 |
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast |
|
552 |
moreover from Cons' have "xs [\<rightarrow>\<^sub>s] ys" by (iprover dest: listrelp_conj1) |
|
553 |
ultimately have "x # xs [\<rightarrow>\<^sub>s] y' # ys" by (rule listrelp.Cons) |
|
554 |
with H show ?thesis by simp |
|
555 |
next |
|
556 |
assume H: "y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'" |
|
557 |
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast |
|
558 |
moreover from H have "xs [\<rightarrow>\<^sub>s] ys'" by (blast intro: Cons') |
|
559 |
ultimately show ?thesis by (rule listrelp.Cons) |
|
560 |
qed |
|
561 |
with ss show ?thesis by simp |
|
562 |
qed |
|
563 |
qed |
|
564 |
||
565 |
lemma lemma4: |
|
566 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
|
567 |
shows "r' \<rightarrow>\<^sub>\<beta> r'' \<Longrightarrow> r \<rightarrow>\<^sub>s r''" using r |
|
568 |
proof (nominal_induct avoiding: r'' rule: sred.strong_induct) |
|
569 |
case (Var rs rs' x) |
|
570 |
then obtain ss where rs: "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss" and r'': "r'' = Var x \<degree>\<degree> ss" |
|
571 |
by (blast dest: head_Var_reduction) |
|
572 |
from Var(1) [simplified] rs have "rs [\<rightarrow>\<^sub>s] ss" by (rule lemma4_aux) |
|
573 |
hence "Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> ss" by (rule sred.Var) |
|
574 |
with r'' show ?case by simp |
|
575 |
next |
|
576 |
case (Abs x ss ss' r r') |
|
63167 | 577 |
from \<open>(Lam [x].r') \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''\<close> show ?case |
27623 | 578 |
proof (cases rule: apps_betasE [where x=x]) |
579 |
case (appL s) |
|
63167 | 580 |
then obtain r''' where s: "s = (Lam [x].r''')" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" using \<open>x \<sharp> r''\<close> |
27623 | 581 |
by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha) |
582 |
from r''' have "r \<rightarrow>\<^sub>s r'''" by (blast intro: Abs) |
|
583 |
moreover from Abs have "ss [\<rightarrow>\<^sub>s] ss'" by (iprover dest: listrelp_conj1) |
|
584 |
ultimately have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r''') \<degree>\<degree> ss'" by (rule better_sred_Abs) |
|
585 |
with appL s show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp |
|
586 |
next |
|
587 |
case (appR rs') |
|
63167 | 588 |
from Abs(6) [simplified] \<open>ss' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs'\<close> |
27623 | 589 |
have "ss [\<rightarrow>\<^sub>s] rs'" by (rule lemma4_aux) |
63167 | 590 |
with \<open>r \<rightarrow>\<^sub>s r'\<close> have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> rs'" by (rule better_sred_Abs) |
27623 | 591 |
with appR show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp |
592 |
next |
|
593 |
case (beta t u' us') |
|
594 |
then have Lam_eq: "(Lam [x].r') = (Lam [x].t)" and ss': "ss' = u' # us'" |
|
595 |
and r'': "r'' = t[x::=u'] \<degree>\<degree> us'" |
|
596 |
by (simp_all add: abs_fresh) |
|
597 |
from Abs(6) ss' obtain u us where |
|
598 |
ss: "ss = u # us" and u: "u \<rightarrow>\<^sub>s u'" and us: "us [\<rightarrow>\<^sub>s] us'" |
|
599 |
by cases (auto dest!: listrelp_conj1) |
|
63167 | 600 |
have "r[x::=u] \<rightarrow>\<^sub>s r'[x::=u']" using \<open>r \<rightarrow>\<^sub>s r'\<close> and u by (rule lemma3) |
27623 | 601 |
with us have "r[x::=u] \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule lemma1') |
602 |
hence "(Lam [x].r) \<degree> u \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule better_sred_Beta) |
|
603 |
with ss r'' Lam_eq show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by (simp add: lam.inject alpha) |
|
604 |
qed |
|
605 |
next |
|
606 |
case (Beta x s ss t r) |
|
607 |
show ?case |
|
608 |
by (rule better_sred_Beta) (rule Beta)+ |
|
609 |
qed |
|
610 |
||
611 |
lemma rtrancl_beta_sred: |
|
612 |
assumes r: "r \<rightarrow>\<^sub>\<beta>\<^sup>* r'" |
|
613 |
shows "r \<rightarrow>\<^sub>s r'" using r |
|
614 |
by induct (iprover intro: refl_sred lemma4)+ |
|
615 |
||
616 |
||
63167 | 617 |
subsection \<open>Terms in normal form\<close> |
27623 | 618 |
|
619 |
lemma listsp_eqvt [eqvt]: |
|
620 |
assumes xs: "listsp p (xs::'a::pt_name list)" |
|
621 |
shows "listsp ((pi::name prm) \<bullet> p) (pi \<bullet> xs)" using xs |
|
80141 | 622 |
by induction (use perm_app in force)+ |
27623 | 623 |
|
624 |
inductive NF :: "lam \<Rightarrow> bool" |
|
625 |
where |
|
626 |
App: "listsp NF ts \<Longrightarrow> NF (Var x \<degree>\<degree> ts)" |
|
627 |
| Abs: "NF t \<Longrightarrow> NF (Lam [x].t)" |
|
628 |
||
629 |
equivariance NF |
|
630 |
nominal_inductive NF |
|
631 |
by (simp add: abs_fresh) |
|
632 |
||
633 |
lemma Abs_NF: |
|
634 |
assumes NF: "NF ((Lam [x].t) \<degree>\<degree> ts)" |
|
635 |
shows "ts = []" using NF |
|
80141 | 636 |
by (metis Abs_eq_apps_conv NF.cases Var_apps_neq_Abs_apps) |
27623 | 637 |
|
63167 | 638 |
text \<open> |
69597 | 639 |
\<^term>\<open>NF\<close> characterizes exactly the terms that are in normal form. |
63167 | 640 |
\<close> |
27623 | 641 |
|
642 |
lemma NF_eq: "NF t = (\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t')" |
|
643 |
proof |
|
644 |
assume H: "NF t" |
|
645 |
show "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'" |
|
646 |
proof |
|
647 |
fix t' |
|
648 |
from H show "\<not> t \<rightarrow>\<^sub>\<beta> t'" |
|
649 |
proof (nominal_induct avoiding: t' rule: NF.strong_induct) |
|
650 |
case (App ts t) |
|
651 |
show ?case |
|
652 |
proof |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
653 |
assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
654 |
then obtain rs where "ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
655 |
by (iprover dest: head_Var_reduction) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
656 |
with App show False |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
657 |
by (induct rs arbitrary: ts) (auto del: in_listspD) |
27623 | 658 |
qed |
659 |
next |
|
660 |
case (Abs t x) |
|
661 |
show ?case |
|
662 |
proof |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
663 |
assume "(Lam [x].t) \<rightarrow>\<^sub>\<beta> t'" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
664 |
then show False using Abs |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
665 |
by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha) |
27623 | 666 |
qed |
667 |
qed |
|
668 |
qed |
|
669 |
next |
|
670 |
assume H: "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'" |
|
671 |
then show "NF t" |
|
672 |
proof (nominal_induct t rule: Apps_lam_induct) |
|
673 |
case (1 n ts) |
|
674 |
then have "\<forall>ts'. \<not> ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ts'" |
|
675 |
by (iprover intro: apps_preserves_betas) |
|
676 |
with 1(1) have "listsp NF ts" |
|
677 |
by (induct ts) (auto simp add: in_listsp_conv_set) |
|
678 |
then show ?case by (rule NF.App) |
|
679 |
next |
|
680 |
case (2 x u ts) |
|
681 |
show ?case |
|
682 |
proof (cases ts) |
|
39459 | 683 |
case Nil thus ?thesis by (metis 2 NF.Abs abs foldl_Nil) |
27623 | 684 |
next |
685 |
case (Cons r rs) |
|
686 |
have "(Lam [x].u) \<degree> r \<rightarrow>\<^sub>\<beta> u[x::=r]" .. |
|
687 |
then have "(Lam [x].u) \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
688 |
by (rule apps_preserves_beta) |
27623 | 689 |
with Cons have "(Lam [x].u) \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29097
diff
changeset
|
690 |
by simp |
27623 | 691 |
with 2 show ?thesis by iprover |
692 |
qed |
|
693 |
qed |
|
694 |
qed |
|
695 |
||
696 |
||
63167 | 697 |
subsection \<open>Leftmost reduction and weakly normalizing terms\<close> |
27623 | 698 |
|
699 |
inductive |
|
700 |
lred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50) |
|
701 |
and lredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>l]" 50) |
|
702 |
where |
|
67399 | 703 |
"s [\<rightarrow>\<^sub>l] t \<equiv> listrelp (\<rightarrow>\<^sub>l) s t" |
27623 | 704 |
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'" |
705 |
| Abs: "r \<rightarrow>\<^sub>l r' \<Longrightarrow> (Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')" |
|
706 |
| Beta: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>l t" |
|
707 |
||
708 |
lemma lred_imp_sred: |
|
709 |
assumes lred: "s \<rightarrow>\<^sub>l t" |
|
710 |
shows "s \<rightarrow>\<^sub>s t" using lred |
|
711 |
proof induct |
|
712 |
case (Var rs rs' x) |
|
713 |
then have "rs [\<rightarrow>\<^sub>s] rs'" |
|
714 |
by induct (iprover intro: listrelp.intros)+ |
|
715 |
then show ?case by (rule sred.Var) |
|
716 |
next |
|
717 |
case (Abs r r' x) |
|
63167 | 718 |
from \<open>r \<rightarrow>\<^sub>s r'\<close> |
27623 | 719 |
have "(Lam [x].r) \<degree>\<degree> [] \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> []" using listrelp.Nil |
720 |
by (rule better_sred_Abs) |
|
721 |
then show ?case by simp |
|
722 |
next |
|
723 |
case (Beta r x s ss t) |
|
63167 | 724 |
from \<open>r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t\<close> |
27623 | 725 |
show ?case by (rule better_sred_Beta) |
726 |
qed |
|
727 |
||
728 |
inductive WN :: "lam \<Rightarrow> bool" |
|
729 |
where |
|
730 |
Var: "listsp WN rs \<Longrightarrow> WN (Var n \<degree>\<degree> rs)" |
|
731 |
| Lambda: "WN r \<Longrightarrow> WN (Lam [x].r)" |
|
732 |
| Beta: "WN ((r[x::=s]) \<degree>\<degree> ss) \<Longrightarrow> WN (((Lam [x].r) \<degree> s) \<degree>\<degree> ss)" |
|
733 |
||
734 |
lemma listrelp_imp_listsp1: |
|
735 |
assumes H: "listrelp (\<lambda>x y. P x) xs ys" |
|
736 |
shows "listsp P xs" using H |
|
737 |
by induct auto |
|
738 |
||
739 |
lemma listrelp_imp_listsp2: |
|
740 |
assumes H: "listrelp (\<lambda>x y. P y) xs ys" |
|
741 |
shows "listsp P ys" using H |
|
742 |
by induct auto |
|
743 |
||
744 |
lemma lemma5: |
|
745 |
assumes lred: "r \<rightarrow>\<^sub>l r'" |
|
746 |
shows "WN r" and "NF r'" using lred |
|
747 |
by induct |
|
748 |
(iprover dest: listrelp_conj1 listrelp_conj2 |
|
749 |
listrelp_imp_listsp1 listrelp_imp_listsp2 intro: WN.intros |
|
750 |
NF.intros)+ |
|
751 |
||
752 |
lemma lemma6: |
|
753 |
assumes wn: "WN r" |
|
754 |
shows "\<exists>r'. r \<rightarrow>\<^sub>l r'" using wn |
|
755 |
proof induct |
|
756 |
case (Var rs n) |
|
757 |
then have "\<exists>rs'. rs [\<rightarrow>\<^sub>l] rs'" |
|
758 |
by induct (iprover intro: listrelp.intros)+ |
|
759 |
then show ?case by (iprover intro: lred.Var) |
|
760 |
qed (iprover intro: lred.intros)+ |
|
761 |
||
762 |
lemma lemma7: |
|
763 |
assumes r: "r \<rightarrow>\<^sub>s r'" |
|
764 |
shows "NF r' \<Longrightarrow> r \<rightarrow>\<^sub>l r'" using r |
|
765 |
proof induct |
|
766 |
case (Var rs rs' x) |
|
63167 | 767 |
from \<open>NF (Var x \<degree>\<degree> rs')\<close> have "listsp NF rs'" |
27623 | 768 |
by cases simp_all |
769 |
with Var(1) have "rs [\<rightarrow>\<^sub>l] rs'" |
|
770 |
proof induct |
|
771 |
case Nil |
|
772 |
show ?case by (rule listrelp.Nil) |
|
773 |
next |
|
39459 | 774 |
case (Cons x y xs ys) |
27623 | 775 |
hence "x \<rightarrow>\<^sub>l y" and "xs [\<rightarrow>\<^sub>l] ys" by (auto del: in_listspD) |
776 |
thus ?case by (rule listrelp.Cons) |
|
777 |
qed |
|
778 |
thus ?case by (rule lred.Var) |
|
779 |
next |
|
780 |
case (Abs x ss ss' r r') |
|
63167 | 781 |
from \<open>NF ((Lam [x].r') \<degree>\<degree> ss')\<close> |
27623 | 782 |
have ss': "ss' = []" by (rule Abs_NF) |
783 |
from Abs(4) have ss: "ss = []" using ss' |
|
784 |
by cases simp_all |
|
785 |
from ss' Abs have "NF (Lam [x].r')" by simp |
|
786 |
hence "NF r'" by (cases rule: NF.strong_cases) (auto simp add: abs_fresh lam.inject alpha) |
|
787 |
with Abs have "r \<rightarrow>\<^sub>l r'" by simp |
|
788 |
hence "(Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')" by (rule lred.Abs) |
|
789 |
with ss ss' show ?case by simp |
|
790 |
next |
|
791 |
case (Beta x s ss t r) |
|
792 |
hence "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t" by simp |
|
793 |
thus ?case by (rule lred.Beta) |
|
794 |
qed |
|
795 |
||
796 |
lemma WN_eq: "WN t = (\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')" |
|
797 |
proof |
|
798 |
assume "WN t" |
|
799 |
then have "\<exists>t'. t \<rightarrow>\<^sub>l t'" by (rule lemma6) |
|
800 |
then obtain t' where t': "t \<rightarrow>\<^sub>l t'" .. |
|
801 |
then have NF: "NF t'" by (rule lemma5) |
|
802 |
from t' have "t \<rightarrow>\<^sub>s t'" by (rule lred_imp_sred) |
|
803 |
then have "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" by (rule lemma2) |
|
804 |
with NF show "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by iprover |
|
805 |
next |
|
806 |
assume "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" |
|
807 |
then obtain t' where t': "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and NF: "NF t'" |
|
808 |
by iprover |
|
809 |
from t' have "t \<rightarrow>\<^sub>s t'" by (rule rtrancl_beta_sred) |
|
810 |
then have "t \<rightarrow>\<^sub>l t'" using NF by (rule lemma7) |
|
811 |
then show "WN t" by (rule lemma5) |
|
812 |
qed |
|
813 |
||
814 |
end |
|
47397
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
haftmann
parents:
46317
diff
changeset
|
815 |