author | wenzelm |
Sat, 17 Jun 2006 19:37:53 +0200 | |
changeset 19912 | 4a3e35fd6e02 |
parent 19496 | 79dbe35c6cba |
permissions | -rw-r--r-- |
18269 | 1 |
(* $Id$ *) |
18106 | 2 |
|
19496 | 3 |
theory Lam_substs |
4 |
imports Iteration |
|
18106 | 5 |
begin |
6 |
||
7 |
constdefs |
|
8 |
depth_Var :: "name \<Rightarrow> nat" |
|
9 |
"depth_Var \<equiv> \<lambda>(a::name). 1" |
|
10 |
||
11 |
depth_App :: "nat \<Rightarrow> nat \<Rightarrow> nat" |
|
12 |
"depth_App \<equiv> \<lambda>n1 n2. (max n1 n2)+1" |
|
13 |
||
14 |
depth_Lam :: "name \<Rightarrow> nat \<Rightarrow> nat" |
|
15 |
"depth_Lam \<equiv> \<lambda>(a::name) n. n+1" |
|
16 |
||
17 |
depth_lam :: "lam \<Rightarrow> nat" |
|
19171 | 18 |
"depth_lam \<equiv> itfun depth_Var depth_App depth_Lam" |
18106 | 19 |
|
20 |
lemma supp_depth_Var: |
|
21 |
shows "((supp depth_Var)::name set)={}" |
|
22 |
apply(simp add: depth_Var_def) |
|
23 |
apply(simp add: supp_def) |
|
24 |
apply(simp add: perm_fun_def) |
|
25 |
apply(simp add: perm_nat_def) |
|
26 |
done |
|
27 |
||
28 |
lemma supp_depth_App: |
|
29 |
shows "((supp depth_App)::name set)={}" |
|
30 |
apply(simp add: depth_App_def) |
|
31 |
apply(simp add: supp_def) |
|
32 |
apply(simp add: perm_fun_def) |
|
33 |
apply(simp add: perm_nat_def) |
|
34 |
done |
|
35 |
||
36 |
lemma supp_depth_Lam: |
|
37 |
shows "((supp depth_Lam)::name set)={}" |
|
38 |
apply(simp add: depth_Lam_def) |
|
39 |
apply(simp add: supp_def) |
|
40 |
apply(simp add: perm_fun_def) |
|
41 |
apply(simp add: perm_nat_def) |
|
42 |
done |
|
43 |
||
44 |
lemma fin_supp_depth: |
|
45 |
shows "finite ((supp (depth_Var,depth_App,depth_Lam))::name set)" |
|
19171 | 46 |
by (finite_guess add: depth_Var_def depth_App_def depth_Lam_def perm_nat_def fs_name1) |
18106 | 47 |
|
48 |
lemma fresh_depth_Lam: |
|
49 |
shows "\<exists>(a::name). a\<sharp>depth_Lam \<and> (\<forall>n. a\<sharp>depth_Lam a n)" |
|
50 |
apply(rule exI) |
|
51 |
apply(rule conjI) |
|
19171 | 52 |
apply(simp add: fresh_def) |
53 |
apply(force simp add: supp_depth_Lam) |
|
18106 | 54 |
apply(unfold fresh_def) |
55 |
apply(simp add: supp_def) |
|
56 |
apply(simp add: perm_nat_def) |
|
57 |
done |
|
58 |
||
59 |
lemma depth_Var: |
|
60 |
shows "depth_lam (Var c) = 1" |
|
61 |
apply(simp add: depth_lam_def) |
|
19171 | 62 |
apply(simp add: itfun_Var[OF fin_supp_depth, OF fresh_depth_Lam]) |
18106 | 63 |
apply(simp add: depth_Var_def) |
64 |
done |
|
65 |
||
66 |
lemma depth_App: |
|
67 |
shows "depth_lam (App t1 t2) = (max (depth_lam t1) (depth_lam t2))+1" |
|
68 |
apply(simp add: depth_lam_def) |
|
19171 | 69 |
apply(simp add: itfun_App[OF fin_supp_depth, OF fresh_depth_Lam]) |
18106 | 70 |
apply(simp add: depth_App_def) |
71 |
done |
|
72 |
||
73 |
lemma depth_Lam: |
|
74 |
shows "depth_lam (Lam [a].t) = (depth_lam t)+1" |
|
75 |
apply(simp add: depth_lam_def) |
|
76 |
apply(rule trans) |
|
19171 | 77 |
apply(rule itfun_Lam[OF fin_supp_depth, OF fresh_depth_Lam]) |
18106 | 78 |
apply(simp add: fresh_def supp_prod supp_depth_Var supp_depth_App supp_depth_Lam) |
79 |
apply(simp add: depth_Lam_def) |
|
80 |
done |
|
81 |
||
19171 | 82 |
text {* substitution *} |
18106 | 83 |
constdefs |
84 |
subst_Var :: "name \<Rightarrow>lam \<Rightarrow> name \<Rightarrow> lam" |
|
85 |
"subst_Var b t \<equiv> \<lambda>a. (if a=b then t else (Var a))" |
|
86 |
||
87 |
subst_App :: "name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam" |
|
88 |
"subst_App b t \<equiv> \<lambda>r1 r2. App r1 r2" |
|
89 |
||
90 |
subst_Lam :: "name \<Rightarrow> lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" |
|
91 |
"subst_Lam b t \<equiv> \<lambda>a r. Lam [a].r" |
|
92 |
||
93 |
subst_lam :: "name \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam" |
|
19171 | 94 |
"subst_lam b t \<equiv> itfun (subst_Var b t) (subst_App b t) (subst_Lam b t)" |
18106 | 95 |
|
96 |
lemma supports_subst_Var: |
|
97 |
shows "((supp (b,t))::name set) supports (subst_Var b t)" |
|
98 |
apply(supports_simp add: subst_Var_def) |
|
99 |
apply(rule impI) |
|
100 |
apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) |
|
101 |
apply(simp add: pt_fresh_fresh[OF pt_name_inst, OF at_name_inst]) |
|
102 |
done |
|
103 |
||
104 |
lemma supports_subst_App: |
|
105 |
shows "((supp (b,t))::name set) supports subst_App b t" |
|
19171 | 106 |
by (supports_simp add: subst_App_def) |
18106 | 107 |
|
108 |
lemma supports_subst_Lam: |
|
109 |
shows "((supp (b,t))::name set) supports subst_Lam b t" |
|
110 |
by (supports_simp add: subst_Lam_def) |
|
111 |
||
112 |
lemma fin_supp_subst: |
|
113 |
shows "finite ((supp (subst_Var b t,subst_App b t,subst_Lam b t))::name set)" |
|
114 |
apply(auto simp add: supp_prod) |
|
115 |
apply(rule supports_finite) |
|
116 |
apply(rule supports_subst_Var) |
|
117 |
apply(simp add: fs_name1) |
|
118 |
apply(rule supports_finite) |
|
119 |
apply(rule supports_subst_App) |
|
120 |
apply(simp add: fs_name1) |
|
121 |
apply(rule supports_finite) |
|
122 |
apply(rule supports_subst_Lam) |
|
123 |
apply(simp add: fs_name1) |
|
124 |
done |
|
125 |
||
126 |
lemma fresh_subst_Lam: |
|
127 |
shows "\<exists>(a::name). a\<sharp>(subst_Lam b t)\<and> (\<forall>y. a\<sharp>(subst_Lam b t) a y)" |
|
128 |
apply(subgoal_tac "\<exists>(c::name). c\<sharp>(b,t)") |
|
129 |
apply(auto) |
|
130 |
apply(rule_tac x="c" in exI) |
|
131 |
apply(auto) |
|
132 |
apply(rule supports_fresh) |
|
133 |
apply(rule supports_subst_Lam) |
|
134 |
apply(simp_all add: fresh_def[symmetric] fs_name1) |
|
135 |
apply(simp add: subst_Lam_def) |
|
136 |
apply(simp add: abs_fresh) |
|
137 |
apply(rule at_exists_fresh[OF at_name_inst]) |
|
138 |
apply(simp add: fs_name1) |
|
139 |
done |
|
140 |
||
141 |
lemma subst_Var: |
|
142 |
shows "subst_lam b t (Var a) = (if a=b then t else (Var a))" |
|
143 |
apply(simp add: subst_lam_def) |
|
19171 | 144 |
apply(auto simp add: itfun_Var[OF fin_supp_subst, OF fresh_subst_Lam]) |
18106 | 145 |
apply(auto simp add: subst_Var_def) |
146 |
done |
|
147 |
||
148 |
lemma subst_App: |
|
149 |
shows "subst_lam b t (App t1 t2) = App (subst_lam b t t1) (subst_lam b t t2)" |
|
150 |
apply(simp add: subst_lam_def) |
|
19171 | 151 |
apply(auto simp add: itfun_App[OF fin_supp_subst, OF fresh_subst_Lam]) |
18106 | 152 |
apply(auto simp add: subst_App_def) |
153 |
done |
|
154 |
||
155 |
lemma subst_Lam: |
|
156 |
assumes a: "a\<sharp>(b,t)" |
|
157 |
shows "subst_lam b t (Lam [a].t1) = Lam [a].(subst_lam b t t1)" |
|
158 |
using a |
|
159 |
apply(simp add: subst_lam_def) |
|
160 |
apply(subgoal_tac "a\<sharp>(subst_Var b t,subst_App b t,subst_Lam b t)") |
|
19171 | 161 |
apply(auto simp add: itfun_Lam[OF fin_supp_subst, OF fresh_subst_Lam]) |
18106 | 162 |
apply(simp (no_asm) add: subst_Lam_def) |
163 |
apply(auto simp add: fresh_prod) |
|
164 |
apply(rule supports_fresh) |
|
165 |
apply(rule supports_subst_Var) |
|
166 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
167 |
apply(rule supports_fresh) |
|
168 |
apply(rule supports_subst_App) |
|
169 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
170 |
apply(rule supports_fresh) |
|
171 |
apply(rule supports_subst_Lam) |
|
172 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
173 |
done |
|
174 |
||
175 |
lemma subst_Lam': |
|
176 |
assumes a: "a\<noteq>b" |
|
177 |
and b: "a\<sharp>t" |
|
178 |
shows "subst_lam b t (Lam [a].t1) = Lam [a].(subst_lam b t t1)" |
|
19171 | 179 |
by (simp add: subst_Lam fresh_prod a b fresh_atm) |
18106 | 180 |
|
18300
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
181 |
lemma subst_Lam'': |
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
182 |
assumes a: "a\<sharp>b" |
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
183 |
and b: "a\<sharp>t" |
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
184 |
shows "subst_lam b t (Lam [a].t1) = Lam [a].(subst_lam b t t1)" |
19171 | 185 |
by (simp add: subst_Lam fresh_prod a b) |
18300
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
186 |
|
18106 | 187 |
consts |
188 |
subst_syn :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 900) |
|
189 |
translations |
|
190 |
"t1[a::=t2]" \<rightleftharpoons> "subst_lam a t2 t1" |
|
191 |
||
192 |
declare subst_Var[simp] |
|
193 |
declare subst_App[simp] |
|
194 |
declare subst_Lam[simp] |
|
195 |
declare subst_Lam'[simp] |
|
18300
ca1ac9e81bc8
added one clause for substitution in the lambda-case and
urbanc
parents:
18298
diff
changeset
|
196 |
declare subst_Lam''[simp] |
18106 | 197 |
|
198 |
lemma subst_eqvt[simp]: |
|
199 |
fixes pi:: "name prm" |
|
200 |
and t1:: "lam" |
|
201 |
and t2:: "lam" |
|
202 |
and a :: "name" |
|
18301
0c5c3b1a700e
fixed the lemma where the new names generated by nominal_induct
urbanc
parents:
18300
diff
changeset
|
203 |
shows "pi\<bullet>(t1[b::=t2]) = (pi\<bullet>t1)[(pi\<bullet>b)::=(pi\<bullet>t2)]" |
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18363
diff
changeset
|
204 |
apply(nominal_induct t1 avoiding: b t2 rule: lam.induct) |
18106 | 205 |
apply(auto simp add: perm_bij fresh_prod fresh_atm) |
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18363
diff
changeset
|
206 |
apply(subgoal_tac "(pi\<bullet>name)\<sharp>(pi\<bullet>b,pi\<bullet>t2)")(*A*) |
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18363
diff
changeset
|
207 |
apply(simp) |
19171 | 208 |
(*A*) |
209 |
apply(simp add: pt_fresh_right[OF pt_name_inst, OF at_name_inst], perm_simp add: fresh_prod fresh_atm) |
|
18106 | 210 |
done |
211 |
||
19171 | 212 |
lemma subst_supp: |
213 |
shows "supp(t1[a::=t2]) \<subseteq> (((supp(t1)-{a})\<union>supp(t2))::name set)" |
|
18659
2ff0ae57431d
changes to make use of the new induction principle proved by
urbanc
parents:
18363
diff
changeset
|
214 |
apply(nominal_induct t1 avoiding: a t2 rule: lam.induct) |
18106 | 215 |
apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp) |
19171 | 216 |
apply(blast)+ |
18106 | 217 |
done |
218 |
||
219 |
(* parallel substitution *) |
|
220 |
||
221 |
consts |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
222 |
"domain" :: "(name\<times>lam) list \<Rightarrow> (name list)" |
18106 | 223 |
primrec |
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
224 |
"domain [] = []" |
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
225 |
"domain (x#\<theta>) = (fst x)#(domain \<theta>)" |
18106 | 226 |
|
227 |
consts |
|
228 |
"apply_sss" :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam" (" _ < _ >" [80,80] 80) |
|
229 |
primrec |
|
230 |
"(x#\<theta>)<a> = (if (a = fst x) then (snd x) else \<theta><a>)" |
|
231 |
||
232 |
||
233 |
lemma apply_sss_eqvt[rule_format]: |
|
234 |
fixes pi::"name prm" |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
235 |
shows "a\<in>set (domain \<theta>) \<longrightarrow> pi\<bullet>(\<theta><a>) = (pi\<bullet>\<theta>)<pi\<bullet>a>" |
18106 | 236 |
apply(induct_tac \<theta>) |
237 |
apply(auto) |
|
238 |
apply(simp add: pt_bij[OF pt_name_inst, OF at_name_inst]) |
|
239 |
done |
|
240 |
||
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
241 |
lemma domain_eqvt[rule_format]: |
18106 | 242 |
fixes pi::"name prm" |
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
243 |
shows "((pi\<bullet>a)\<in>set (domain \<theta>)) = (a\<in>set (domain ((rev pi)\<bullet>\<theta>)))" |
18106 | 244 |
apply(induct_tac \<theta>) |
245 |
apply(auto) |
|
246 |
apply(simp_all add: pt_rev_pi[OF pt_name_inst, OF at_name_inst]) |
|
247 |
apply(simp_all add: pt_pi_rev[OF pt_name_inst, OF at_name_inst]) |
|
248 |
done |
|
249 |
||
250 |
constdefs |
|
251 |
psubst_Var :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam" |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
252 |
"psubst_Var \<theta> \<equiv> \<lambda>a. (if a\<in>set (domain \<theta>) then \<theta><a> else (Var a))" |
18106 | 253 |
|
254 |
psubst_App :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam \<Rightarrow> lam" |
|
255 |
"psubst_App \<theta> \<equiv> \<lambda>r1 r2. App r1 r2" |
|
256 |
||
257 |
psubst_Lam :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" |
|
258 |
"psubst_Lam \<theta> \<equiv> \<lambda>a r. Lam [a].r" |
|
259 |
||
260 |
psubst_lam :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" |
|
19171 | 261 |
"psubst_lam \<theta> \<equiv> itfun (psubst_Var \<theta>) (psubst_App \<theta>) (psubst_Lam \<theta>)" |
18106 | 262 |
|
263 |
lemma supports_psubst_Var: |
|
264 |
shows "((supp \<theta>)::name set) supports (psubst_Var \<theta>)" |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
265 |
by (supports_simp add: psubst_Var_def apply_sss_eqvt domain_eqvt) |
18106 | 266 |
|
267 |
lemma supports_psubst_App: |
|
268 |
shows "((supp \<theta>)::name set) supports psubst_App \<theta>" |
|
269 |
by (supports_simp add: psubst_App_def) |
|
270 |
||
271 |
lemma supports_psubst_Lam: |
|
272 |
shows "((supp \<theta>)::name set) supports psubst_Lam \<theta>" |
|
273 |
by (supports_simp add: psubst_Lam_def) |
|
274 |
||
275 |
lemma fin_supp_psubst: |
|
276 |
shows "finite ((supp (psubst_Var \<theta>,psubst_App \<theta>,psubst_Lam \<theta>))::name set)" |
|
277 |
apply(auto simp add: supp_prod) |
|
278 |
apply(rule supports_finite) |
|
279 |
apply(rule supports_psubst_Var) |
|
280 |
apply(simp add: fs_name1) |
|
281 |
apply(rule supports_finite) |
|
282 |
apply(rule supports_psubst_App) |
|
283 |
apply(simp add: fs_name1) |
|
284 |
apply(rule supports_finite) |
|
285 |
apply(rule supports_psubst_Lam) |
|
286 |
apply(simp add: fs_name1) |
|
287 |
done |
|
288 |
||
289 |
lemma fresh_psubst_Lam: |
|
290 |
shows "\<exists>(a::name). a\<sharp>(psubst_Lam \<theta>)\<and> (\<forall>y. a\<sharp>(psubst_Lam \<theta>) a y)" |
|
291 |
apply(subgoal_tac "\<exists>(c::name). c\<sharp>\<theta>") |
|
292 |
apply(auto) |
|
293 |
apply(rule_tac x="c" in exI) |
|
294 |
apply(auto) |
|
295 |
apply(rule supports_fresh) |
|
296 |
apply(rule supports_psubst_Lam) |
|
297 |
apply(simp_all add: fresh_def[symmetric] fs_name1) |
|
298 |
apply(simp add: psubst_Lam_def) |
|
299 |
apply(simp add: abs_fresh) |
|
300 |
apply(rule at_exists_fresh[OF at_name_inst]) |
|
301 |
apply(simp add: fs_name1) |
|
302 |
done |
|
303 |
||
304 |
lemma psubst_Var: |
|
18263
7f75925498da
cleaned up all examples so that they work with the
urbanc
parents:
18106
diff
changeset
|
305 |
shows "psubst_lam \<theta> (Var a) = (if a\<in>set (domain \<theta>) then \<theta><a> else (Var a))" |
18106 | 306 |
apply(simp add: psubst_lam_def) |
19171 | 307 |
apply(auto simp add: itfun_Var[OF fin_supp_psubst, OF fresh_psubst_Lam]) |
18106 | 308 |
apply(auto simp add: psubst_Var_def) |
309 |
done |
|
310 |
||
311 |
lemma psubst_App: |
|
312 |
shows "psubst_lam \<theta> (App t1 t2) = App (psubst_lam \<theta> t1) (psubst_lam \<theta> t2)" |
|
313 |
apply(simp add: psubst_lam_def) |
|
19171 | 314 |
apply(auto simp add: itfun_App[OF fin_supp_psubst, OF fresh_psubst_Lam]) |
18106 | 315 |
apply(auto simp add: psubst_App_def) |
316 |
done |
|
317 |
||
318 |
lemma psubst_Lam: |
|
319 |
assumes a: "a\<sharp>\<theta>" |
|
320 |
shows "psubst_lam \<theta> (Lam [a].t1) = Lam [a].(psubst_lam \<theta> t1)" |
|
321 |
using a |
|
322 |
apply(simp add: psubst_lam_def) |
|
323 |
apply(subgoal_tac "a\<sharp>(psubst_Var \<theta>,psubst_App \<theta>,psubst_Lam \<theta>)") |
|
19171 | 324 |
apply(auto simp add: itfun_Lam[OF fin_supp_psubst, OF fresh_psubst_Lam]) |
18106 | 325 |
apply(simp (no_asm) add: psubst_Lam_def) |
326 |
apply(auto simp add: fresh_prod) |
|
327 |
apply(rule supports_fresh) |
|
328 |
apply(rule supports_psubst_Var) |
|
329 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
330 |
apply(rule supports_fresh) |
|
331 |
apply(rule supports_psubst_App) |
|
332 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
333 |
apply(rule supports_fresh) |
|
334 |
apply(rule supports_psubst_Lam) |
|
335 |
apply(simp_all add: fs_name1 fresh_def[symmetric] fresh_prod) |
|
336 |
done |
|
337 |
||
338 |
declare psubst_Var[simp] |
|
339 |
declare psubst_App[simp] |
|
340 |
declare psubst_Lam[simp] |
|
341 |
||
342 |
consts |
|
343 |
psubst_syn :: "lam \<Rightarrow> (name\<times>lam) list \<Rightarrow> lam" ("_[<_>]" [100,100] 900) |
|
344 |
translations |
|
345 |
"t[<\<theta>>]" \<rightleftharpoons> "psubst_lam \<theta> t" |
|
346 |
||
347 |
end |