author | wenzelm |
Mon, 29 Jul 2002 00:57:16 +0200 | |
changeset 13428 | 99e52e78eb65 |
parent 13422 | af9bc8d87a75 |
child 13429 | 2232810416fc |
permissions | -rw-r--r-- |
13363 | 1 |
header{*Separation for Facts About Recursion*} |
13348 | 2 |
|
13363 | 3 |
theory Rec_Separation = Separation + Datatype_absolute: |
13348 | 4 |
|
5 |
text{*This theory proves all instances needed for locales @{text |
|
6 |
"M_trancl"}, @{text "M_wfrank"} and @{text "M_datatypes"}*} |
|
7 |
||
13363 | 8 |
lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> j<i" |
13428 | 9 |
by simp |
13363 | 10 |
|
13348 | 11 |
subsection{*The Locale @{text "M_trancl"}*} |
12 |
||
13 |
subsubsection{*Separation for Reflexive/Transitive Closure*} |
|
14 |
||
15 |
text{*First, The Defining Formula*} |
|
16 |
||
17 |
(* "rtran_closure_mem(M,A,r,p) == |
|
13428 | 18 |
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. |
13348 | 19 |
omega(M,nnat) & n\<in>nnat & successor(M,n,n') & |
20 |
(\<exists>f[M]. typed_function(M,n',A,f) & |
|
13428 | 21 |
(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) & |
22 |
fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & |
|
23 |
(\<forall>j[M]. j\<in>n --> |
|
24 |
(\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. |
|
25 |
fun_apply(M,f,j,fj) & successor(M,j,sj) & |
|
26 |
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*) |
|
13348 | 27 |
constdefs rtran_closure_mem_fm :: "[i,i,i]=>i" |
13428 | 28 |
"rtran_closure_mem_fm(A,r,p) == |
13348 | 29 |
Exists(Exists(Exists( |
30 |
And(omega_fm(2), |
|
31 |
And(Member(1,2), |
|
32 |
And(succ_fm(1,0), |
|
33 |
Exists(And(typed_function_fm(1, A#+4, 0), |
|
13428 | 34 |
And(Exists(Exists(Exists( |
35 |
And(pair_fm(2,1,p#+7), |
|
36 |
And(empty_fm(0), |
|
37 |
And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))), |
|
38 |
Forall(Implies(Member(0,3), |
|
39 |
Exists(Exists(Exists(Exists( |
|
40 |
And(fun_apply_fm(5,4,3), |
|
41 |
And(succ_fm(4,2), |
|
42 |
And(fun_apply_fm(5,2,1), |
|
43 |
And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))" |
|
13348 | 44 |
|
45 |
||
46 |
lemma rtran_closure_mem_type [TC]: |
|
47 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> rtran_closure_mem_fm(x,y,z) \<in> formula" |
|
13428 | 48 |
by (simp add: rtran_closure_mem_fm_def) |
13348 | 49 |
|
50 |
lemma arity_rtran_closure_mem_fm [simp]: |
|
13428 | 51 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] |
13348 | 52 |
==> arity(rtran_closure_mem_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)" |
13428 | 53 |
by (simp add: rtran_closure_mem_fm_def succ_Un_distrib [symmetric] Un_ac) |
13348 | 54 |
|
55 |
lemma sats_rtran_closure_mem_fm [simp]: |
|
56 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
13428 | 57 |
==> sats(A, rtran_closure_mem_fm(x,y,z), env) <-> |
13348 | 58 |
rtran_closure_mem(**A, nth(x,env), nth(y,env), nth(z,env))" |
59 |
by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def) |
|
60 |
||
61 |
lemma rtran_closure_mem_iff_sats: |
|
13428 | 62 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13348 | 63 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
64 |
==> rtran_closure_mem(**A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)" |
|
65 |
by (simp add: sats_rtran_closure_mem_fm) |
|
66 |
||
67 |
theorem rtran_closure_mem_reflection: |
|
13428 | 68 |
"REFLECTS[\<lambda>x. rtran_closure_mem(L,f(x),g(x),h(x)), |
13348 | 69 |
\<lambda>i x. rtran_closure_mem(**Lset(i),f(x),g(x),h(x))]" |
70 |
apply (simp only: rtran_closure_mem_def setclass_simps) |
|
13428 | 71 |
apply (intro FOL_reflections function_reflections fun_plus_reflections) |
13348 | 72 |
done |
73 |
||
74 |
text{*Separation for @{term "rtrancl(r)"}.*} |
|
75 |
lemma rtrancl_separation: |
|
76 |
"[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))" |
|
13428 | 77 |
apply (rule separation_CollectI) |
78 |
apply (rule_tac A="{r,A,z}" in subset_LsetE, blast ) |
|
13348 | 79 |
apply (rule ReflectsE [OF rtran_closure_mem_reflection], assumption) |
13428 | 80 |
apply (drule subset_Lset_ltD, assumption) |
13348 | 81 |
apply (erule reflection_imp_L_separation) |
82 |
apply (simp_all add: lt_Ord2) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
83 |
apply (rule DPow_LsetI) |
13348 | 84 |
apply (rename_tac u) |
85 |
apply (rule_tac env = "[u,r,A]" in rtran_closure_mem_iff_sats) |
|
86 |
apply (rule sep_rules | simp)+ |
|
87 |
done |
|
88 |
||
89 |
||
90 |
subsubsection{*Reflexive/Transitive Closure, Internalized*} |
|
91 |
||
13428 | 92 |
(* "rtran_closure(M,r,s) == |
13348 | 93 |
\<forall>A[M]. is_field(M,r,A) --> |
13428 | 94 |
(\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *) |
13348 | 95 |
constdefs rtran_closure_fm :: "[i,i]=>i" |
13428 | 96 |
"rtran_closure_fm(r,s) == |
13348 | 97 |
Forall(Implies(field_fm(succ(r),0), |
98 |
Forall(Iff(Member(0,succ(succ(s))), |
|
99 |
rtran_closure_mem_fm(1,succ(succ(r)),0)))))" |
|
100 |
||
101 |
lemma rtran_closure_type [TC]: |
|
102 |
"[| x \<in> nat; y \<in> nat |] ==> rtran_closure_fm(x,y) \<in> formula" |
|
13428 | 103 |
by (simp add: rtran_closure_fm_def) |
13348 | 104 |
|
105 |
lemma arity_rtran_closure_fm [simp]: |
|
13428 | 106 |
"[| x \<in> nat; y \<in> nat |] |
13348 | 107 |
==> arity(rtran_closure_fm(x,y)) = succ(x) \<union> succ(y)" |
108 |
by (simp add: rtran_closure_fm_def succ_Un_distrib [symmetric] Un_ac) |
|
109 |
||
110 |
lemma sats_rtran_closure_fm [simp]: |
|
111 |
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|] |
|
13428 | 112 |
==> sats(A, rtran_closure_fm(x,y), env) <-> |
13348 | 113 |
rtran_closure(**A, nth(x,env), nth(y,env))" |
114 |
by (simp add: rtran_closure_fm_def rtran_closure_def) |
|
115 |
||
116 |
lemma rtran_closure_iff_sats: |
|
13428 | 117 |
"[| nth(i,env) = x; nth(j,env) = y; |
13348 | 118 |
i \<in> nat; j \<in> nat; env \<in> list(A)|] |
119 |
==> rtran_closure(**A, x, y) <-> sats(A, rtran_closure_fm(i,j), env)" |
|
120 |
by simp |
|
121 |
||
122 |
theorem rtran_closure_reflection: |
|
13428 | 123 |
"REFLECTS[\<lambda>x. rtran_closure(L,f(x),g(x)), |
13348 | 124 |
\<lambda>i x. rtran_closure(**Lset(i),f(x),g(x))]" |
125 |
apply (simp only: rtran_closure_def setclass_simps) |
|
126 |
apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection) |
|
127 |
done |
|
128 |
||
129 |
||
130 |
subsubsection{*Transitive Closure of a Relation, Internalized*} |
|
131 |
||
132 |
(* "tran_closure(M,r,t) == |
|
133 |
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *) |
|
134 |
constdefs tran_closure_fm :: "[i,i]=>i" |
|
13428 | 135 |
"tran_closure_fm(r,s) == |
13348 | 136 |
Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" |
137 |
||
138 |
lemma tran_closure_type [TC]: |
|
139 |
"[| x \<in> nat; y \<in> nat |] ==> tran_closure_fm(x,y) \<in> formula" |
|
13428 | 140 |
by (simp add: tran_closure_fm_def) |
13348 | 141 |
|
142 |
lemma arity_tran_closure_fm [simp]: |
|
13428 | 143 |
"[| x \<in> nat; y \<in> nat |] |
13348 | 144 |
==> arity(tran_closure_fm(x,y)) = succ(x) \<union> succ(y)" |
145 |
by (simp add: tran_closure_fm_def succ_Un_distrib [symmetric] Un_ac) |
|
146 |
||
147 |
lemma sats_tran_closure_fm [simp]: |
|
148 |
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|] |
|
13428 | 149 |
==> sats(A, tran_closure_fm(x,y), env) <-> |
13348 | 150 |
tran_closure(**A, nth(x,env), nth(y,env))" |
151 |
by (simp add: tran_closure_fm_def tran_closure_def) |
|
152 |
||
153 |
lemma tran_closure_iff_sats: |
|
13428 | 154 |
"[| nth(i,env) = x; nth(j,env) = y; |
13348 | 155 |
i \<in> nat; j \<in> nat; env \<in> list(A)|] |
156 |
==> tran_closure(**A, x, y) <-> sats(A, tran_closure_fm(i,j), env)" |
|
157 |
by simp |
|
158 |
||
159 |
theorem tran_closure_reflection: |
|
13428 | 160 |
"REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)), |
13348 | 161 |
\<lambda>i x. tran_closure(**Lset(i),f(x),g(x))]" |
162 |
apply (simp only: tran_closure_def setclass_simps) |
|
13428 | 163 |
apply (intro FOL_reflections function_reflections |
13348 | 164 |
rtran_closure_reflection composition_reflection) |
165 |
done |
|
166 |
||
167 |
||
168 |
subsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*} |
|
169 |
||
170 |
lemma wellfounded_trancl_reflects: |
|
13428 | 171 |
"REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L]. |
172 |
w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp, |
|
173 |
\<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i). |
|
13348 | 174 |
w \<in> Z & pair(**Lset(i),w,x,wx) & tran_closure(**Lset(i),r,rp) & |
175 |
wx \<in> rp]" |
|
13428 | 176 |
by (intro FOL_reflections function_reflections fun_plus_reflections |
13348 | 177 |
tran_closure_reflection) |
178 |
||
179 |
||
180 |
lemma wellfounded_trancl_separation: |
|
13428 | 181 |
"[| L(r); L(Z) |] ==> |
182 |
separation (L, \<lambda>x. |
|
183 |
\<exists>w[L]. \<exists>wx[L]. \<exists>rp[L]. |
|
184 |
w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)" |
|
185 |
apply (rule separation_CollectI) |
|
186 |
apply (rule_tac A="{r,Z,z}" in subset_LsetE, blast ) |
|
13348 | 187 |
apply (rule ReflectsE [OF wellfounded_trancl_reflects], assumption) |
13428 | 188 |
apply (drule subset_Lset_ltD, assumption) |
13348 | 189 |
apply (erule reflection_imp_L_separation) |
190 |
apply (simp_all add: lt_Ord2) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
191 |
apply (rule DPow_LsetI) |
13428 | 192 |
apply (rename_tac u) |
13348 | 193 |
apply (rule bex_iff_sats conj_iff_sats)+ |
13428 | 194 |
apply (rule_tac env = "[w,u,r,Z]" in mem_iff_sats) |
13348 | 195 |
apply (rule sep_rules tran_closure_iff_sats | simp)+ |
196 |
done |
|
197 |
||
13363 | 198 |
|
199 |
subsubsection{*Instantiating the locale @{text M_trancl}*} |
|
13428 | 200 |
|
201 |
theorem M_trancl_axioms_L: "M_trancl_axioms(L)" |
|
202 |
apply (rule M_trancl_axioms.intro) |
|
203 |
apply (assumption | rule |
|
204 |
rtrancl_separation wellfounded_trancl_separation)+ |
|
205 |
done |
|
13363 | 206 |
|
13428 | 207 |
theorem M_trancl_L: "PROP M_trancl(L)" |
208 |
apply (rule M_trancl.intro) |
|
209 |
apply (rule M_triv_axioms_L) |
|
210 |
apply (rule M_axioms_axioms_L) |
|
211 |
apply (rule M_trancl_axioms_L) |
|
212 |
done |
|
13363 | 213 |
|
13428 | 214 |
lemmas iterates_abs = M_trancl.iterates_abs [OF M_trancl_L] |
215 |
and rtran_closure_rtrancl = M_trancl.rtran_closure_rtrancl [OF M_trancl_L] |
|
216 |
and rtrancl_closed = M_trancl.rtrancl_closed [OF M_trancl_L] |
|
217 |
and rtrancl_abs = M_trancl.rtrancl_abs [OF M_trancl_L] |
|
218 |
and trancl_closed = M_trancl.trancl_closed [OF M_trancl_L] |
|
219 |
and trancl_abs = M_trancl.trancl_abs [OF M_trancl_L] |
|
220 |
and wellfounded_on_trancl = M_trancl.wellfounded_on_trancl [OF M_trancl_L] |
|
221 |
and wellfounded_trancl = M_trancl.wellfounded_trancl [OF M_trancl_L] |
|
222 |
and wfrec_relativize = M_trancl.wfrec_relativize [OF M_trancl_L] |
|
223 |
and trans_wfrec_relativize = M_trancl.trans_wfrec_relativize [OF M_trancl_L] |
|
224 |
and trans_wfrec_abs = M_trancl.trans_wfrec_abs [OF M_trancl_L] |
|
225 |
and trans_eq_pair_wfrec_iff = M_trancl.trans_eq_pair_wfrec_iff [OF M_trancl_L] |
|
226 |
and eq_pair_wfrec_iff = M_trancl.eq_pair_wfrec_iff [OF M_trancl_L] |
|
13363 | 227 |
|
228 |
declare rtrancl_closed [intro,simp] |
|
229 |
declare rtrancl_abs [simp] |
|
230 |
declare trancl_closed [intro,simp] |
|
231 |
declare trancl_abs [simp] |
|
232 |
||
233 |
||
13348 | 234 |
subsection{*Well-Founded Recursion!*} |
235 |
||
13352 | 236 |
(* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" |
13428 | 237 |
"M_is_recfun(M,MH,r,a,f) == |
238 |
\<forall>z[M]. z \<in> f <-> |
|
13348 | 239 |
5 4 3 2 1 0 |
13428 | 240 |
(\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. |
241 |
pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) & |
|
13348 | 242 |
pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) & |
243 |
xa \<in> r & MH(x, f_r_sx, y))" |
|
244 |
*) |
|
245 |
||
246 |
constdefs is_recfun_fm :: "[[i,i,i]=>i, i, i, i]=>i" |
|
13428 | 247 |
"is_recfun_fm(p,r,a,f) == |
13348 | 248 |
Forall(Iff(Member(0,succ(f)), |
249 |
Exists(Exists(Exists(Exists(Exists(Exists( |
|
250 |
And(pair_fm(5,4,6), |
|
251 |
And(pair_fm(5,a#+7,3), |
|
252 |
And(upair_fm(5,5,2), |
|
253 |
And(pre_image_fm(r#+7,2,1), |
|
254 |
And(restriction_fm(f#+7,1,0), |
|
255 |
And(Member(3,r#+7), p(5,0,4)))))))))))))))" |
|
256 |
||
257 |
||
258 |
lemma is_recfun_type_0: |
|
13428 | 259 |
"[| !!x y z. [| x \<in> nat; y \<in> nat; z \<in> nat |] ==> p(x,y,z) \<in> formula; |
260 |
x \<in> nat; y \<in> nat; z \<in> nat |] |
|
13348 | 261 |
==> is_recfun_fm(p,x,y,z) \<in> formula" |
262 |
apply (unfold is_recfun_fm_def) |
|
263 |
(*FIXME: FIND OUT why simp loops!*) |
|
264 |
apply typecheck |
|
13428 | 265 |
by simp |
13348 | 266 |
|
267 |
lemma is_recfun_type [TC]: |
|
13428 | 268 |
"[| p(5,0,4) \<in> formula; |
269 |
x \<in> nat; y \<in> nat; z \<in> nat |] |
|
13348 | 270 |
==> is_recfun_fm(p,x,y,z) \<in> formula" |
13428 | 271 |
by (simp add: is_recfun_fm_def) |
13348 | 272 |
|
273 |
lemma arity_is_recfun_fm [simp]: |
|
13428 | 274 |
"[| arity(p(5,0,4)) le 8; x \<in> nat; y \<in> nat; z \<in> nat |] |
13348 | 275 |
==> arity(is_recfun_fm(p,x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)" |
13428 | 276 |
apply (frule lt_nat_in_nat, simp) |
277 |
apply (simp add: is_recfun_fm_def succ_Un_distrib [symmetric] ) |
|
278 |
apply (subst subset_Un_iff2 [of "arity(p(5,0,4))", THEN iffD1]) |
|
279 |
apply (rule le_imp_subset) |
|
280 |
apply (erule le_trans, simp) |
|
281 |
apply (simp add: succ_Un_distrib [symmetric] Un_ac) |
|
13348 | 282 |
done |
283 |
||
284 |
lemma sats_is_recfun_fm: |
|
13428 | 285 |
assumes MH_iff_sats: |
286 |
"!!x y z env. |
|
287 |
[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
288 |
==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> sats(A, p(x,y,z), env)" |
|
289 |
shows |
|
13348 | 290 |
"[|x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
13428 | 291 |
==> sats(A, is_recfun_fm(p,x,y,z), env) <-> |
13352 | 292 |
M_is_recfun(**A, MH, nth(x,env), nth(y,env), nth(z,env))" |
13348 | 293 |
by (simp add: is_recfun_fm_def M_is_recfun_def MH_iff_sats [THEN iff_sym]) |
294 |
||
295 |
lemma is_recfun_iff_sats: |
|
296 |
"[| (!!x y z env. [| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
297 |
==> MH(nth(x,env), nth(y,env), nth(z,env)) <-> |
|
298 |
sats(A, p(x,y,z), env)); |
|
13428 | 299 |
nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13348 | 300 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
13428 | 301 |
==> M_is_recfun(**A, MH, x, y, z) <-> sats(A, is_recfun_fm(p,i,j,k), env)" |
13348 | 302 |
by (simp add: sats_is_recfun_fm [of A MH]) |
303 |
||
304 |
theorem is_recfun_reflection: |
|
305 |
assumes MH_reflection: |
|
13428 | 306 |
"!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)), |
13348 | 307 |
\<lambda>i x. MH(**Lset(i), f(x), g(x), h(x))]" |
13428 | 308 |
shows "REFLECTS[\<lambda>x. M_is_recfun(L, MH(L), f(x), g(x), h(x)), |
13352 | 309 |
\<lambda>i x. M_is_recfun(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]" |
13348 | 310 |
apply (simp (no_asm_use) only: M_is_recfun_def setclass_simps) |
13428 | 311 |
apply (intro FOL_reflections function_reflections |
312 |
restriction_reflection MH_reflection) |
|
13348 | 313 |
done |
314 |
||
13363 | 315 |
text{*Currently, @{text sats}-theorems for higher-order operators don't seem |
316 |
useful. Reflection theorems do work, though. This one avoids the repetition |
|
317 |
of the @{text MH}-term.*} |
|
318 |
theorem is_wfrec_reflection: |
|
319 |
assumes MH_reflection: |
|
13428 | 320 |
"!!f g h. REFLECTS[\<lambda>x. MH(L, f(x), g(x), h(x)), |
13363 | 321 |
\<lambda>i x. MH(**Lset(i), f(x), g(x), h(x))]" |
13428 | 322 |
shows "REFLECTS[\<lambda>x. is_wfrec(L, MH(L), f(x), g(x), h(x)), |
13363 | 323 |
\<lambda>i x. is_wfrec(**Lset(i), MH(**Lset(i)), f(x), g(x), h(x))]" |
324 |
apply (simp (no_asm_use) only: is_wfrec_def setclass_simps) |
|
13428 | 325 |
apply (intro FOL_reflections MH_reflection is_recfun_reflection) |
13363 | 326 |
done |
327 |
||
328 |
subsection{*The Locale @{text "M_wfrank"}*} |
|
329 |
||
330 |
subsubsection{*Separation for @{term "wfrank"}*} |
|
13348 | 331 |
|
332 |
lemma wfrank_Reflects: |
|
333 |
"REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) --> |
|
13352 | 334 |
~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)), |
13348 | 335 |
\<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) --> |
13428 | 336 |
~ (\<exists>f \<in> Lset(i). |
337 |
M_is_recfun(**Lset(i), %x f y. is_range(**Lset(i),f,y), |
|
13352 | 338 |
rplus, x, f))]" |
13428 | 339 |
by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection) |
13348 | 340 |
|
341 |
lemma wfrank_separation: |
|
342 |
"L(r) ==> |
|
343 |
separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) --> |
|
13352 | 344 |
~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))" |
13428 | 345 |
apply (rule separation_CollectI) |
346 |
apply (rule_tac A="{r,z}" in subset_LsetE, blast ) |
|
13348 | 347 |
apply (rule ReflectsE [OF wfrank_Reflects], assumption) |
13428 | 348 |
apply (drule subset_Lset_ltD, assumption) |
13348 | 349 |
apply (erule reflection_imp_L_separation) |
350 |
apply (simp_all add: lt_Ord2, clarify) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
351 |
apply (rule DPow_LsetI) |
13428 | 352 |
apply (rename_tac u) |
13348 | 353 |
apply (rule ball_iff_sats imp_iff_sats)+ |
354 |
apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats) |
|
355 |
apply (rule sep_rules is_recfun_iff_sats | simp)+ |
|
356 |
done |
|
357 |
||
358 |
||
13363 | 359 |
subsubsection{*Replacement for @{term "wfrank"}*} |
13348 | 360 |
|
361 |
lemma wfrank_replacement_Reflects: |
|
13428 | 362 |
"REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A & |
13348 | 363 |
(\<forall>rplus[L]. tran_closure(L,r,rplus) --> |
13428 | 364 |
(\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) & |
13352 | 365 |
M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & |
13348 | 366 |
is_range(L,f,y))), |
13428 | 367 |
\<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A & |
13348 | 368 |
(\<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) --> |
13428 | 369 |
(\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(**Lset(i),x,y,z) & |
13352 | 370 |
M_is_recfun(**Lset(i), %x f y. is_range(**Lset(i),f,y), rplus, x, f) & |
13348 | 371 |
is_range(**Lset(i),f,y)))]" |
372 |
by (intro FOL_reflections function_reflections fun_plus_reflections |
|
373 |
is_recfun_reflection tran_closure_reflection) |
|
374 |
||
375 |
||
376 |
lemma wfrank_strong_replacement: |
|
377 |
"L(r) ==> |
|
13428 | 378 |
strong_replacement(L, \<lambda>x z. |
13348 | 379 |
\<forall>rplus[L]. tran_closure(L,r,rplus) --> |
13428 | 380 |
(\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z) & |
13352 | 381 |
M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) & |
13348 | 382 |
is_range(L,f,y)))" |
13428 | 383 |
apply (rule strong_replacementI) |
13348 | 384 |
apply (rule rallI) |
13428 | 385 |
apply (rename_tac B) |
386 |
apply (rule separation_CollectI) |
|
387 |
apply (rule_tac A="{B,r,z}" in subset_LsetE, blast ) |
|
13348 | 388 |
apply (rule ReflectsE [OF wfrank_replacement_Reflects], assumption) |
13428 | 389 |
apply (drule subset_Lset_ltD, assumption) |
13348 | 390 |
apply (erule reflection_imp_L_separation) |
391 |
apply (simp_all add: lt_Ord2) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
392 |
apply (rule DPow_LsetI) |
13428 | 393 |
apply (rename_tac u) |
13348 | 394 |
apply (rule bex_iff_sats ball_iff_sats conj_iff_sats)+ |
13428 | 395 |
apply (rule_tac env = "[x,u,B,r]" in mem_iff_sats) |
13348 | 396 |
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+ |
397 |
done |
|
398 |
||
399 |
||
13363 | 400 |
subsubsection{*Separation for Proving @{text Ord_wfrank_range}*} |
13348 | 401 |
|
402 |
lemma Ord_wfrank_Reflects: |
|
13428 | 403 |
"REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) --> |
404 |
~ (\<forall>f[L]. \<forall>rangef[L]. |
|
13348 | 405 |
is_range(L,f,rangef) --> |
13352 | 406 |
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) --> |
13348 | 407 |
ordinal(L,rangef)), |
13428 | 408 |
\<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(**Lset(i),r,rplus) --> |
409 |
~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i). |
|
13348 | 410 |
is_range(**Lset(i),f,rangef) --> |
13428 | 411 |
M_is_recfun(**Lset(i), \<lambda>x f y. is_range(**Lset(i),f,y), |
13352 | 412 |
rplus, x, f) --> |
13348 | 413 |
ordinal(**Lset(i),rangef))]" |
13428 | 414 |
by (intro FOL_reflections function_reflections is_recfun_reflection |
13348 | 415 |
tran_closure_reflection ordinal_reflection) |
416 |
||
417 |
lemma Ord_wfrank_separation: |
|
418 |
"L(r) ==> |
|
419 |
separation (L, \<lambda>x. |
|
13428 | 420 |
\<forall>rplus[L]. tran_closure(L,r,rplus) --> |
421 |
~ (\<forall>f[L]. \<forall>rangef[L]. |
|
13348 | 422 |
is_range(L,f,rangef) --> |
13352 | 423 |
M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) --> |
13428 | 424 |
ordinal(L,rangef)))" |
425 |
apply (rule separation_CollectI) |
|
426 |
apply (rule_tac A="{r,z}" in subset_LsetE, blast ) |
|
13348 | 427 |
apply (rule ReflectsE [OF Ord_wfrank_Reflects], assumption) |
13428 | 428 |
apply (drule subset_Lset_ltD, assumption) |
13348 | 429 |
apply (erule reflection_imp_L_separation) |
430 |
apply (simp_all add: lt_Ord2, clarify) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
431 |
apply (rule DPow_LsetI) |
13428 | 432 |
apply (rename_tac u) |
13348 | 433 |
apply (rule ball_iff_sats imp_iff_sats)+ |
434 |
apply (rule_tac env="[rplus,u,r]" in tran_closure_iff_sats) |
|
435 |
apply (rule sep_rules is_recfun_iff_sats | simp)+ |
|
436 |
done |
|
437 |
||
438 |
||
13363 | 439 |
subsubsection{*Instantiating the locale @{text M_wfrank}*} |
13428 | 440 |
|
441 |
theorem M_wfrank_axioms_L: "M_wfrank_axioms(L)" |
|
442 |
apply (rule M_wfrank_axioms.intro) |
|
443 |
apply (assumption | rule |
|
444 |
wfrank_separation wfrank_strong_replacement Ord_wfrank_separation)+ |
|
445 |
done |
|
13363 | 446 |
|
13428 | 447 |
theorem M_wfrank_L: "PROP M_wfrank(L)" |
448 |
apply (rule M_wfrank.intro) |
|
449 |
apply (rule M_triv_axioms_L) |
|
450 |
apply (rule M_axioms_axioms_L) |
|
451 |
apply (rule M_trancl_axioms_L) |
|
452 |
apply (rule M_wfrank_axioms_L) |
|
453 |
done |
|
13363 | 454 |
|
13428 | 455 |
lemmas iterates_closed = M_wfrank.iterates_closed [OF M_wfrank_L] |
456 |
and exists_wfrank = M_wfrank.exists_wfrank [OF M_wfrank_L] |
|
457 |
and M_wellfoundedrank = M_wfrank.M_wellfoundedrank [OF M_wfrank_L] |
|
458 |
and Ord_wfrank_range = M_wfrank.Ord_wfrank_range [OF M_wfrank_L] |
|
459 |
and Ord_range_wellfoundedrank = M_wfrank.Ord_range_wellfoundedrank [OF M_wfrank_L] |
|
460 |
and function_wellfoundedrank = M_wfrank.function_wellfoundedrank [OF M_wfrank_L] |
|
461 |
and domain_wellfoundedrank = M_wfrank.domain_wellfoundedrank [OF M_wfrank_L] |
|
462 |
and wellfoundedrank_type = M_wfrank.wellfoundedrank_type [OF M_wfrank_L] |
|
463 |
and Ord_wellfoundedrank = M_wfrank.Ord_wellfoundedrank [OF M_wfrank_L] |
|
464 |
and wellfoundedrank_eq = M_wfrank.wellfoundedrank_eq [OF M_wfrank_L] |
|
465 |
and wellfoundedrank_lt = M_wfrank.wellfoundedrank_lt [OF M_wfrank_L] |
|
466 |
and wellfounded_imp_subset_rvimage = M_wfrank.wellfounded_imp_subset_rvimage [OF M_wfrank_L] |
|
467 |
and wellfounded_imp_wf = M_wfrank.wellfounded_imp_wf [OF M_wfrank_L] |
|
468 |
and wellfounded_on_imp_wf_on = M_wfrank.wellfounded_on_imp_wf_on [OF M_wfrank_L] |
|
469 |
and wf_abs = M_wfrank.wf_abs [OF M_wfrank_L] |
|
470 |
and wf_on_abs = M_wfrank.wf_on_abs [OF M_wfrank_L] |
|
471 |
and wfrec_replacement_iff = M_wfrank.wfrec_replacement_iff [OF M_wfrank_L] |
|
472 |
and trans_wfrec_closed = M_wfrank.trans_wfrec_closed [OF M_wfrank_L] |
|
473 |
and wfrec_closed = M_wfrank.wfrec_closed [OF M_wfrank_L] |
|
13363 | 474 |
|
475 |
declare iterates_closed [intro,simp] |
|
476 |
declare Ord_wfrank_range [rule_format] |
|
477 |
declare wf_abs [simp] |
|
478 |
declare wf_on_abs [simp] |
|
479 |
||
480 |
||
481 |
subsection{*For Datatypes*} |
|
482 |
||
483 |
subsubsection{*Binary Products, Internalized*} |
|
484 |
||
485 |
constdefs cartprod_fm :: "[i,i,i]=>i" |
|
13428 | 486 |
(* "cartprod(M,A,B,z) == |
487 |
\<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *) |
|
488 |
"cartprod_fm(A,B,z) == |
|
13363 | 489 |
Forall(Iff(Member(0,succ(z)), |
490 |
Exists(And(Member(0,succ(succ(A))), |
|
491 |
Exists(And(Member(0,succ(succ(succ(B)))), |
|
492 |
pair_fm(1,0,2)))))))" |
|
493 |
||
494 |
lemma cartprod_type [TC]: |
|
495 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> cartprod_fm(x,y,z) \<in> formula" |
|
13428 | 496 |
by (simp add: cartprod_fm_def) |
13363 | 497 |
|
498 |
lemma arity_cartprod_fm [simp]: |
|
13428 | 499 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] |
13363 | 500 |
==> arity(cartprod_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)" |
13428 | 501 |
by (simp add: cartprod_fm_def succ_Un_distrib [symmetric] Un_ac) |
13363 | 502 |
|
503 |
lemma sats_cartprod_fm [simp]: |
|
504 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
13428 | 505 |
==> sats(A, cartprod_fm(x,y,z), env) <-> |
13363 | 506 |
cartprod(**A, nth(x,env), nth(y,env), nth(z,env))" |
507 |
by (simp add: cartprod_fm_def cartprod_def) |
|
508 |
||
509 |
lemma cartprod_iff_sats: |
|
13428 | 510 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13363 | 511 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
512 |
==> cartprod(**A, x, y, z) <-> sats(A, cartprod_fm(i,j,k), env)" |
|
513 |
by (simp add: sats_cartprod_fm) |
|
514 |
||
515 |
theorem cartprod_reflection: |
|
13428 | 516 |
"REFLECTS[\<lambda>x. cartprod(L,f(x),g(x),h(x)), |
13363 | 517 |
\<lambda>i x. cartprod(**Lset(i),f(x),g(x),h(x))]" |
518 |
apply (simp only: cartprod_def setclass_simps) |
|
13428 | 519 |
apply (intro FOL_reflections pair_reflection) |
13363 | 520 |
done |
521 |
||
522 |
||
523 |
subsubsection{*Binary Sums, Internalized*} |
|
524 |
||
13428 | 525 |
(* "is_sum(M,A,B,Z) == |
526 |
\<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M]. |
|
13363 | 527 |
3 2 1 0 |
528 |
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) & |
|
529 |
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *) |
|
530 |
constdefs sum_fm :: "[i,i,i]=>i" |
|
13428 | 531 |
"sum_fm(A,B,Z) == |
13363 | 532 |
Exists(Exists(Exists(Exists( |
13428 | 533 |
And(number1_fm(2), |
13363 | 534 |
And(cartprod_fm(2,A#+4,3), |
535 |
And(upair_fm(2,2,1), |
|
536 |
And(cartprod_fm(1,B#+4,0), union_fm(3,0,Z#+4)))))))))" |
|
537 |
||
538 |
lemma sum_type [TC]: |
|
539 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> sum_fm(x,y,z) \<in> formula" |
|
13428 | 540 |
by (simp add: sum_fm_def) |
13363 | 541 |
|
542 |
lemma arity_sum_fm [simp]: |
|
13428 | 543 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] |
13363 | 544 |
==> arity(sum_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)" |
13428 | 545 |
by (simp add: sum_fm_def succ_Un_distrib [symmetric] Un_ac) |
13363 | 546 |
|
547 |
lemma sats_sum_fm [simp]: |
|
548 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
13428 | 549 |
==> sats(A, sum_fm(x,y,z), env) <-> |
13363 | 550 |
is_sum(**A, nth(x,env), nth(y,env), nth(z,env))" |
551 |
by (simp add: sum_fm_def is_sum_def) |
|
552 |
||
553 |
lemma sum_iff_sats: |
|
13428 | 554 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13363 | 555 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
556 |
==> is_sum(**A, x, y, z) <-> sats(A, sum_fm(i,j,k), env)" |
|
557 |
by simp |
|
558 |
||
559 |
theorem sum_reflection: |
|
13428 | 560 |
"REFLECTS[\<lambda>x. is_sum(L,f(x),g(x),h(x)), |
13363 | 561 |
\<lambda>i x. is_sum(**Lset(i),f(x),g(x),h(x))]" |
562 |
apply (simp only: is_sum_def setclass_simps) |
|
13428 | 563 |
apply (intro FOL_reflections function_reflections cartprod_reflection) |
13363 | 564 |
done |
565 |
||
566 |
||
567 |
subsubsection{*The Operator @{term quasinat}*} |
|
568 |
||
569 |
(* "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *) |
|
570 |
constdefs quasinat_fm :: "i=>i" |
|
571 |
"quasinat_fm(z) == Or(empty_fm(z), Exists(succ_fm(0,succ(z))))" |
|
572 |
||
573 |
lemma quasinat_type [TC]: |
|
574 |
"x \<in> nat ==> quasinat_fm(x) \<in> formula" |
|
13428 | 575 |
by (simp add: quasinat_fm_def) |
13363 | 576 |
|
577 |
lemma arity_quasinat_fm [simp]: |
|
578 |
"x \<in> nat ==> arity(quasinat_fm(x)) = succ(x)" |
|
13428 | 579 |
by (simp add: quasinat_fm_def succ_Un_distrib [symmetric] Un_ac) |
13363 | 580 |
|
581 |
lemma sats_quasinat_fm [simp]: |
|
582 |
"[| x \<in> nat; env \<in> list(A)|] |
|
583 |
==> sats(A, quasinat_fm(x), env) <-> is_quasinat(**A, nth(x,env))" |
|
584 |
by (simp add: quasinat_fm_def is_quasinat_def) |
|
585 |
||
586 |
lemma quasinat_iff_sats: |
|
13428 | 587 |
"[| nth(i,env) = x; nth(j,env) = y; |
13363 | 588 |
i \<in> nat; env \<in> list(A)|] |
589 |
==> is_quasinat(**A, x) <-> sats(A, quasinat_fm(i), env)" |
|
590 |
by simp |
|
591 |
||
592 |
theorem quasinat_reflection: |
|
13428 | 593 |
"REFLECTS[\<lambda>x. is_quasinat(L,f(x)), |
13363 | 594 |
\<lambda>i x. is_quasinat(**Lset(i),f(x))]" |
595 |
apply (simp only: is_quasinat_def setclass_simps) |
|
13428 | 596 |
apply (intro FOL_reflections function_reflections) |
13363 | 597 |
done |
598 |
||
599 |
||
600 |
subsubsection{*The Operator @{term is_nat_case}*} |
|
601 |
||
602 |
(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" |
|
13428 | 603 |
"is_nat_case(M, a, is_b, k, z) == |
13363 | 604 |
(empty(M,k) --> z=a) & |
605 |
(\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) & |
|
606 |
(is_quasinat(M,k) | empty(M,z))" *) |
|
607 |
text{*The formula @{term is_b} has free variables 1 and 0.*} |
|
608 |
constdefs is_nat_case_fm :: "[i, [i,i]=>i, i, i]=>i" |
|
13428 | 609 |
"is_nat_case_fm(a,is_b,k,z) == |
13363 | 610 |
And(Implies(empty_fm(k), Equal(z,a)), |
13428 | 611 |
And(Forall(Implies(succ_fm(0,succ(k)), |
13363 | 612 |
Forall(Implies(Equal(0,succ(succ(z))), is_b(1,0))))), |
613 |
Or(quasinat_fm(k), empty_fm(z))))" |
|
614 |
||
615 |
lemma is_nat_case_type [TC]: |
|
13428 | 616 |
"[| is_b(1,0) \<in> formula; |
617 |
x \<in> nat; y \<in> nat; z \<in> nat |] |
|
13363 | 618 |
==> is_nat_case_fm(x,is_b,y,z) \<in> formula" |
13428 | 619 |
by (simp add: is_nat_case_fm_def) |
13363 | 620 |
|
621 |
lemma arity_is_nat_case_fm [simp]: |
|
13428 | 622 |
"[| is_b(1,0) \<in> formula; x \<in> nat; y \<in> nat; z \<in> nat |] |
623 |
==> arity(is_nat_case_fm(x,is_b,y,z)) = |
|
624 |
succ(x) \<union> succ(y) \<union> succ(z) \<union> (arity(is_b(1,0)) #- 2)" |
|
625 |
apply (subgoal_tac "arity(is_b(1,0)) \<in> nat") |
|
13363 | 626 |
apply typecheck |
627 |
(*FIXME: could nat_diff_split work?*) |
|
628 |
apply (auto simp add: diff_def raw_diff_succ is_nat_case_fm_def nat_imp_quasinat |
|
629 |
succ_Un_distrib [symmetric] Un_ac |
|
13428 | 630 |
split: split_nat_case) |
13363 | 631 |
done |
632 |
||
633 |
lemma sats_is_nat_case_fm: |
|
13428 | 634 |
assumes is_b_iff_sats: |
635 |
"!!a b. [| a \<in> A; b \<in> A|] |
|
13363 | 636 |
==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env)))" |
13428 | 637 |
shows |
13363 | 638 |
"[|x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|] |
13428 | 639 |
==> sats(A, is_nat_case_fm(x,p,y,z), env) <-> |
13363 | 640 |
is_nat_case(**A, nth(x,env), is_b, nth(y,env), nth(z,env))" |
13428 | 641 |
apply (frule lt_length_in_nat, assumption) |
13363 | 642 |
apply (simp add: is_nat_case_fm_def is_nat_case_def is_b_iff_sats [THEN iff_sym]) |
643 |
done |
|
644 |
||
645 |
lemma is_nat_case_iff_sats: |
|
13428 | 646 |
"[| (!!a b. [| a \<in> A; b \<in> A|] |
13363 | 647 |
==> is_b(a,b) <-> sats(A, p(1,0), Cons(b, Cons(a,env)))); |
13428 | 648 |
nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13363 | 649 |
i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|] |
13428 | 650 |
==> is_nat_case(**A, x, is_b, y, z) <-> sats(A, is_nat_case_fm(i,p,j,k), env)" |
13363 | 651 |
by (simp add: sats_is_nat_case_fm [of A is_b]) |
652 |
||
653 |
||
654 |
text{*The second argument of @{term is_b} gives it direct access to @{term x}, |
|
13428 | 655 |
which is essential for handling free variable references. Without this |
13363 | 656 |
argument, we cannot prove reflection for @{term iterates_MH}.*} |
657 |
theorem is_nat_case_reflection: |
|
658 |
assumes is_b_reflection: |
|
13428 | 659 |
"!!h f g. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x)), |
13363 | 660 |
\<lambda>i x. is_b(**Lset(i), h(x), f(x), g(x))]" |
13428 | 661 |
shows "REFLECTS[\<lambda>x. is_nat_case(L, f(x), is_b(L,x), g(x), h(x)), |
13363 | 662 |
\<lambda>i x. is_nat_case(**Lset(i), f(x), is_b(**Lset(i), x), g(x), h(x))]" |
663 |
apply (simp (no_asm_use) only: is_nat_case_def setclass_simps) |
|
13428 | 664 |
apply (intro FOL_reflections function_reflections |
665 |
restriction_reflection is_b_reflection quasinat_reflection) |
|
13363 | 666 |
done |
667 |
||
668 |
||
669 |
||
670 |
subsection{*The Operator @{term iterates_MH}, Needed for Iteration*} |
|
671 |
||
672 |
(* iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" |
|
673 |
"iterates_MH(M,isF,v,n,g,z) == |
|
674 |
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u), |
|
675 |
n, z)" *) |
|
676 |
constdefs iterates_MH_fm :: "[[i,i]=>i, i, i, i, i]=>i" |
|
13428 | 677 |
"iterates_MH_fm(isF,v,n,g,z) == |
678 |
is_nat_case_fm(v, |
|
679 |
\<lambda>m u. Exists(And(fun_apply_fm(succ(succ(succ(g))),succ(m),0), |
|
680 |
Forall(Implies(Equal(0,succ(succ(u))), isF(1,0))))), |
|
13363 | 681 |
n, z)" |
682 |
||
683 |
lemma iterates_MH_type [TC]: |
|
13428 | 684 |
"[| p(1,0) \<in> formula; |
685 |
v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |] |
|
13363 | 686 |
==> iterates_MH_fm(p,v,x,y,z) \<in> formula" |
13428 | 687 |
by (simp add: iterates_MH_fm_def) |
13363 | 688 |
|
689 |
||
690 |
lemma arity_iterates_MH_fm [simp]: |
|
13428 | 691 |
"[| p(1,0) \<in> formula; |
692 |
v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |] |
|
693 |
==> arity(iterates_MH_fm(p,v,x,y,z)) = |
|
13363 | 694 |
succ(v) \<union> succ(x) \<union> succ(y) \<union> succ(z) \<union> (arity(p(1,0)) #- 4)" |
695 |
apply (subgoal_tac "arity(p(1,0)) \<in> nat") |
|
696 |
apply typecheck |
|
697 |
apply (simp add: nat_imp_quasinat iterates_MH_fm_def Un_ac |
|
698 |
split: split_nat_case, clarify) |
|
699 |
apply (rename_tac i j) |
|
13428 | 700 |
apply (drule eq_succ_imp_eq_m1, simp) |
13363 | 701 |
apply (drule eq_succ_imp_eq_m1, simp) |
702 |
apply (simp add: diff_Un_distrib succ_Un_distrib Un_ac diff_diff_left) |
|
703 |
done |
|
704 |
||
705 |
lemma sats_iterates_MH_fm: |
|
13428 | 706 |
assumes is_F_iff_sats: |
707 |
"!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|] |
|
13363 | 708 |
==> is_F(a,b) <-> |
709 |
sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env)))))" |
|
13428 | 710 |
shows |
13363 | 711 |
"[|v \<in> nat; x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A)|] |
13428 | 712 |
==> sats(A, iterates_MH_fm(p,v,x,y,z), env) <-> |
13363 | 713 |
iterates_MH(**A, is_F, nth(v,env), nth(x,env), nth(y,env), nth(z,env))" |
13428 | 714 |
by (simp add: iterates_MH_fm_def iterates_MH_def sats_is_nat_case_fm |
13363 | 715 |
is_F_iff_sats [symmetric]) |
716 |
||
717 |
lemma iterates_MH_iff_sats: |
|
13428 | 718 |
"[| (!!a b c d. [| a \<in> A; b \<in> A; c \<in> A; d \<in> A|] |
13363 | 719 |
==> is_F(a,b) <-> |
720 |
sats(A, p(1,0), Cons(b, Cons(a, Cons(c, Cons(d,env)))))); |
|
13428 | 721 |
nth(i',env) = v; nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13363 | 722 |
i' \<in> nat; i \<in> nat; j \<in> nat; k < length(env); env \<in> list(A)|] |
13428 | 723 |
==> iterates_MH(**A, is_F, v, x, y, z) <-> |
13363 | 724 |
sats(A, iterates_MH_fm(p,i',i,j,k), env)" |
13428 | 725 |
apply (rule iff_sym) |
726 |
apply (rule iff_trans) |
|
727 |
apply (rule sats_iterates_MH_fm [of A is_F], blast, simp_all) |
|
13363 | 728 |
done |
729 |
||
730 |
theorem iterates_MH_reflection: |
|
731 |
assumes p_reflection: |
|
13428 | 732 |
"!!f g h. REFLECTS[\<lambda>x. p(L, f(x), g(x)), |
13363 | 733 |
\<lambda>i x. p(**Lset(i), f(x), g(x))]" |
13428 | 734 |
shows "REFLECTS[\<lambda>x. iterates_MH(L, p(L), e(x), f(x), g(x), h(x)), |
13363 | 735 |
\<lambda>i x. iterates_MH(**Lset(i), p(**Lset(i)), e(x), f(x), g(x), h(x))]" |
736 |
apply (simp (no_asm_use) only: iterates_MH_def) |
|
737 |
txt{*Must be careful: simplifying with @{text setclass_simps} above would |
|
738 |
change @{text "\<exists>gm[**Lset(i)]"} into @{text "\<exists>gm \<in> Lset(i)"}, when |
|
739 |
it would no longer match rule @{text is_nat_case_reflection}. *} |
|
13428 | 740 |
apply (rule is_nat_case_reflection) |
13363 | 741 |
apply (simp (no_asm_use) only: setclass_simps) |
742 |
apply (intro FOL_reflections function_reflections is_nat_case_reflection |
|
13428 | 743 |
restriction_reflection p_reflection) |
13363 | 744 |
done |
745 |
||
746 |
||
747 |
||
13428 | 748 |
subsection{*@{term L} is Closed Under the Operator @{term list}*} |
13363 | 749 |
|
13386 | 750 |
subsubsection{*The List Functor, Internalized*} |
751 |
||
752 |
constdefs list_functor_fm :: "[i,i,i]=>i" |
|
13428 | 753 |
(* "is_list_functor(M,A,X,Z) == |
754 |
\<exists>n1[M]. \<exists>AX[M]. |
|
13386 | 755 |
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *) |
13428 | 756 |
"list_functor_fm(A,X,Z) == |
13386 | 757 |
Exists(Exists( |
13428 | 758 |
And(number1_fm(1), |
13386 | 759 |
And(cartprod_fm(A#+2,X#+2,0), sum_fm(1,0,Z#+2)))))" |
760 |
||
761 |
lemma list_functor_type [TC]: |
|
762 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> list_functor_fm(x,y,z) \<in> formula" |
|
13428 | 763 |
by (simp add: list_functor_fm_def) |
13386 | 764 |
|
765 |
lemma arity_list_functor_fm [simp]: |
|
13428 | 766 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] |
13386 | 767 |
==> arity(list_functor_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)" |
13428 | 768 |
by (simp add: list_functor_fm_def succ_Un_distrib [symmetric] Un_ac) |
13386 | 769 |
|
770 |
lemma sats_list_functor_fm [simp]: |
|
771 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
|
13428 | 772 |
==> sats(A, list_functor_fm(x,y,z), env) <-> |
13386 | 773 |
is_list_functor(**A, nth(x,env), nth(y,env), nth(z,env))" |
774 |
by (simp add: list_functor_fm_def is_list_functor_def) |
|
775 |
||
776 |
lemma list_functor_iff_sats: |
|
13428 | 777 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13386 | 778 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
779 |
==> is_list_functor(**A, x, y, z) <-> sats(A, list_functor_fm(i,j,k), env)" |
|
780 |
by simp |
|
781 |
||
782 |
theorem list_functor_reflection: |
|
13428 | 783 |
"REFLECTS[\<lambda>x. is_list_functor(L,f(x),g(x),h(x)), |
13386 | 784 |
\<lambda>i x. is_list_functor(**Lset(i),f(x),g(x),h(x))]" |
785 |
apply (simp only: is_list_functor_def setclass_simps) |
|
786 |
apply (intro FOL_reflections number1_reflection |
|
13428 | 787 |
cartprod_reflection sum_reflection) |
13386 | 788 |
done |
789 |
||
790 |
||
791 |
subsubsection{*Instances of Replacement for Lists*} |
|
792 |
||
13363 | 793 |
lemma list_replacement1_Reflects: |
794 |
"REFLECTS |
|
795 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and> |
|
796 |
is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)), |
|
797 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
|
13428 | 798 |
is_wfrec(**Lset(i), |
799 |
iterates_MH(**Lset(i), |
|
13363 | 800 |
is_list_functor(**Lset(i), A), 0), memsn, u, y))]" |
13428 | 801 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
802 |
iterates_MH_reflection list_functor_reflection) |
|
13363 | 803 |
|
13428 | 804 |
lemma list_replacement1: |
13363 | 805 |
"L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)" |
806 |
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) |
|
13428 | 807 |
apply (rule strong_replacementI) |
13363 | 808 |
apply (rule rallI) |
13428 | 809 |
apply (rename_tac B) |
810 |
apply (rule separation_CollectI) |
|
811 |
apply (insert nonempty) |
|
812 |
apply (subgoal_tac "L(Memrel(succ(n)))") |
|
813 |
apply (rule_tac A="{B,A,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) |
|
13363 | 814 |
apply (rule ReflectsE [OF list_replacement1_Reflects], assumption) |
13428 | 815 |
apply (drule subset_Lset_ltD, assumption) |
13363 | 816 |
apply (erule reflection_imp_L_separation) |
13386 | 817 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
13428 | 818 |
apply (elim conjE) |
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
819 |
apply (rule DPow_LsetI) |
13428 | 820 |
apply (rename_tac v) |
13363 | 821 |
apply (rule bex_iff_sats conj_iff_sats)+ |
822 |
apply (rule_tac env = "[u,v,A,n,B,0,Memrel(succ(n))]" in mem_iff_sats) |
|
823 |
apply (rule sep_rules | simp)+ |
|
824 |
txt{*Can't get sat rules to work for higher-order operators, so just expand them!*} |
|
825 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
|
826 |
apply (rule sep_rules list_functor_iff_sats quasinat_iff_sats | simp)+ |
|
827 |
done |
|
828 |
||
829 |
||
830 |
lemma list_replacement2_Reflects: |
|
831 |
"REFLECTS |
|
832 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and> |
|
833 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L, u, sn) \<and> membership(L, sn, msn) \<and> |
|
834 |
is_wfrec (L, iterates_MH (L, is_list_functor(L, A), 0), |
|
835 |
msn, u, x)), |
|
836 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and> |
|
13428 | 837 |
(\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i). |
13363 | 838 |
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and> |
13428 | 839 |
is_wfrec (**Lset(i), |
13363 | 840 |
iterates_MH (**Lset(i), is_list_functor(**Lset(i), A), 0), |
841 |
msn, u, x))]" |
|
13428 | 842 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
843 |
iterates_MH_reflection list_functor_reflection) |
|
13363 | 844 |
|
845 |
||
13428 | 846 |
lemma list_replacement2: |
847 |
"L(A) ==> strong_replacement(L, |
|
848 |
\<lambda>n y. n\<in>nat & |
|
13363 | 849 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) & |
13428 | 850 |
is_wfrec(L, iterates_MH(L,is_list_functor(L,A), 0), |
13363 | 851 |
msn, n, y)))" |
13428 | 852 |
apply (rule strong_replacementI) |
13363 | 853 |
apply (rule rallI) |
13428 | 854 |
apply (rename_tac B) |
855 |
apply (rule separation_CollectI) |
|
856 |
apply (insert nonempty) |
|
857 |
apply (rule_tac A="{A,B,z,0,nat}" in subset_LsetE) |
|
858 |
apply (blast intro: L_nat) |
|
13363 | 859 |
apply (rule ReflectsE [OF list_replacement2_Reflects], assumption) |
13428 | 860 |
apply (drule subset_Lset_ltD, assumption) |
13363 | 861 |
apply (erule reflection_imp_L_separation) |
862 |
apply (simp_all add: lt_Ord2) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
863 |
apply (rule DPow_LsetI) |
13428 | 864 |
apply (rename_tac v) |
13363 | 865 |
apply (rule bex_iff_sats conj_iff_sats)+ |
866 |
apply (rule_tac env = "[u,v,A,B,0,nat]" in mem_iff_sats) |
|
867 |
apply (rule sep_rules | simp)+ |
|
868 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
|
869 |
apply (rule sep_rules list_functor_iff_sats quasinat_iff_sats | simp)+ |
|
870 |
done |
|
871 |
||
13386 | 872 |
|
13428 | 873 |
subsection{*@{term L} is Closed Under the Operator @{term formula}*} |
13386 | 874 |
|
875 |
subsubsection{*The Formula Functor, Internalized*} |
|
876 |
||
877 |
constdefs formula_functor_fm :: "[i,i]=>i" |
|
13428 | 878 |
(* "is_formula_functor(M,X,Z) == |
879 |
\<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M]. |
|
13398
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
880 |
4 3 2 1 0 |
13428 | 881 |
omega(M,nat') & cartprod(M,nat',nat',natnat) & |
13386 | 882 |
is_sum(M,natnat,natnat,natnatsum) & |
13428 | 883 |
cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & |
884 |
is_sum(M,natnatsum,X3,Z)" *) |
|
885 |
"formula_functor_fm(X,Z) == |
|
13398
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
886 |
Exists(Exists(Exists(Exists(Exists( |
13428 | 887 |
And(omega_fm(4), |
13398
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
888 |
And(cartprod_fm(4,4,3), |
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
889 |
And(sum_fm(3,3,2), |
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
890 |
And(cartprod_fm(X#+5,X#+5,1), |
1cadd412da48
Towards relativization and absoluteness of formula_rec
paulson
parents:
13395
diff
changeset
|
891 |
And(sum_fm(1,X#+5,0), sum_fm(2,0,Z#+5)))))))))))" |
13386 | 892 |
|
893 |
lemma formula_functor_type [TC]: |
|
894 |
"[| x \<in> nat; y \<in> nat |] ==> formula_functor_fm(x,y) \<in> formula" |
|
13428 | 895 |
by (simp add: formula_functor_fm_def) |
13386 | 896 |
|
897 |
lemma sats_formula_functor_fm [simp]: |
|
898 |
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|] |
|
13428 | 899 |
==> sats(A, formula_functor_fm(x,y), env) <-> |
13386 | 900 |
is_formula_functor(**A, nth(x,env), nth(y,env))" |
901 |
by (simp add: formula_functor_fm_def is_formula_functor_def) |
|
902 |
||
903 |
lemma formula_functor_iff_sats: |
|
13428 | 904 |
"[| nth(i,env) = x; nth(j,env) = y; |
13386 | 905 |
i \<in> nat; j \<in> nat; env \<in> list(A)|] |
906 |
==> is_formula_functor(**A, x, y) <-> sats(A, formula_functor_fm(i,j), env)" |
|
907 |
by simp |
|
908 |
||
909 |
theorem formula_functor_reflection: |
|
13428 | 910 |
"REFLECTS[\<lambda>x. is_formula_functor(L,f(x),g(x)), |
13386 | 911 |
\<lambda>i x. is_formula_functor(**Lset(i),f(x),g(x))]" |
912 |
apply (simp only: is_formula_functor_def setclass_simps) |
|
913 |
apply (intro FOL_reflections omega_reflection |
|
13428 | 914 |
cartprod_reflection sum_reflection) |
13386 | 915 |
done |
916 |
||
917 |
subsubsection{*Instances of Replacement for Formulas*} |
|
918 |
||
919 |
lemma formula_replacement1_Reflects: |
|
920 |
"REFLECTS |
|
921 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and> |
|
922 |
is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)), |
|
923 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
|
13428 | 924 |
is_wfrec(**Lset(i), |
925 |
iterates_MH(**Lset(i), |
|
13386 | 926 |
is_formula_functor(**Lset(i)), 0), memsn, u, y))]" |
13428 | 927 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
928 |
iterates_MH_reflection formula_functor_reflection) |
|
13386 | 929 |
|
13428 | 930 |
lemma formula_replacement1: |
13386 | 931 |
"iterates_replacement(L, is_formula_functor(L), 0)" |
932 |
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) |
|
13428 | 933 |
apply (rule strong_replacementI) |
13386 | 934 |
apply (rule rallI) |
13428 | 935 |
apply (rename_tac B) |
936 |
apply (rule separation_CollectI) |
|
937 |
apply (insert nonempty) |
|
938 |
apply (subgoal_tac "L(Memrel(succ(n)))") |
|
939 |
apply (rule_tac A="{B,n,z,0,Memrel(succ(n))}" in subset_LsetE, blast ) |
|
13386 | 940 |
apply (rule ReflectsE [OF formula_replacement1_Reflects], assumption) |
13428 | 941 |
apply (drule subset_Lset_ltD, assumption) |
13386 | 942 |
apply (erule reflection_imp_L_separation) |
943 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
|
944 |
apply (rule DPow_LsetI) |
|
13428 | 945 |
apply (rename_tac v) |
13386 | 946 |
apply (rule bex_iff_sats conj_iff_sats)+ |
947 |
apply (rule_tac env = "[u,v,n,B,0,Memrel(succ(n))]" in mem_iff_sats) |
|
948 |
apply (rule sep_rules | simp)+ |
|
949 |
txt{*Can't get sat rules to work for higher-order operators, so just expand them!*} |
|
950 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
|
951 |
apply (rule sep_rules formula_functor_iff_sats quasinat_iff_sats | simp)+ |
|
952 |
txt{*SLOW: like 40 seconds!*} |
|
953 |
done |
|
954 |
||
955 |
lemma formula_replacement2_Reflects: |
|
956 |
"REFLECTS |
|
957 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and> |
|
958 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L, u, sn) \<and> membership(L, sn, msn) \<and> |
|
959 |
is_wfrec (L, iterates_MH (L, is_formula_functor(L), 0), |
|
960 |
msn, u, x)), |
|
961 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and> |
|
13428 | 962 |
(\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i). |
13386 | 963 |
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and> |
13428 | 964 |
is_wfrec (**Lset(i), |
13386 | 965 |
iterates_MH (**Lset(i), is_formula_functor(**Lset(i)), 0), |
966 |
msn, u, x))]" |
|
13428 | 967 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
968 |
iterates_MH_reflection formula_functor_reflection) |
|
13386 | 969 |
|
970 |
||
13428 | 971 |
lemma formula_replacement2: |
972 |
"strong_replacement(L, |
|
973 |
\<lambda>n y. n\<in>nat & |
|
13386 | 974 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) & |
13428 | 975 |
is_wfrec(L, iterates_MH(L,is_formula_functor(L), 0), |
13386 | 976 |
msn, n, y)))" |
13428 | 977 |
apply (rule strong_replacementI) |
13386 | 978 |
apply (rule rallI) |
13428 | 979 |
apply (rename_tac B) |
980 |
apply (rule separation_CollectI) |
|
981 |
apply (insert nonempty) |
|
982 |
apply (rule_tac A="{B,z,0,nat}" in subset_LsetE) |
|
983 |
apply (blast intro: L_nat) |
|
13386 | 984 |
apply (rule ReflectsE [OF formula_replacement2_Reflects], assumption) |
13428 | 985 |
apply (drule subset_Lset_ltD, assumption) |
13386 | 986 |
apply (erule reflection_imp_L_separation) |
987 |
apply (simp_all add: lt_Ord2) |
|
988 |
apply (rule DPow_LsetI) |
|
13428 | 989 |
apply (rename_tac v) |
13386 | 990 |
apply (rule bex_iff_sats conj_iff_sats)+ |
991 |
apply (rule_tac env = "[u,v,B,0,nat]" in mem_iff_sats) |
|
992 |
apply (rule sep_rules | simp)+ |
|
993 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
|
994 |
apply (rule sep_rules formula_functor_iff_sats quasinat_iff_sats | simp)+ |
|
995 |
done |
|
996 |
||
997 |
text{*NB The proofs for type @{term formula} are virtually identical to those |
|
998 |
for @{term "list(A)"}. It was a cut-and-paste job! *} |
|
999 |
||
13387 | 1000 |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1001 |
subsection{*Internalized Forms of Data Structuring Operators*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1002 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1003 |
subsubsection{*The Formula @{term is_Inl}, Internalized*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1004 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1005 |
(* is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1006 |
constdefs Inl_fm :: "[i,i]=>i" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1007 |
"Inl_fm(a,z) == Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1008 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1009 |
lemma Inl_type [TC]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1010 |
"[| x \<in> nat; z \<in> nat |] ==> Inl_fm(x,z) \<in> formula" |
13428 | 1011 |
by (simp add: Inl_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1012 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1013 |
lemma sats_Inl_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1014 |
"[| x \<in> nat; z \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1015 |
==> sats(A, Inl_fm(x,z), env) <-> is_Inl(**A, nth(x,env), nth(z,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1016 |
by (simp add: Inl_fm_def is_Inl_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1017 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1018 |
lemma Inl_iff_sats: |
13428 | 1019 |
"[| nth(i,env) = x; nth(k,env) = z; |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1020 |
i \<in> nat; k \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1021 |
==> is_Inl(**A, x, z) <-> sats(A, Inl_fm(i,k), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1022 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1023 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1024 |
theorem Inl_reflection: |
13428 | 1025 |
"REFLECTS[\<lambda>x. is_Inl(L,f(x),h(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1026 |
\<lambda>i x. is_Inl(**Lset(i),f(x),h(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1027 |
apply (simp only: is_Inl_def setclass_simps) |
13428 | 1028 |
apply (intro FOL_reflections function_reflections) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1029 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1030 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1031 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1032 |
subsubsection{*The Formula @{term is_Inr}, Internalized*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1033 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1034 |
(* is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1035 |
constdefs Inr_fm :: "[i,i]=>i" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1036 |
"Inr_fm(a,z) == Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1037 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1038 |
lemma Inr_type [TC]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1039 |
"[| x \<in> nat; z \<in> nat |] ==> Inr_fm(x,z) \<in> formula" |
13428 | 1040 |
by (simp add: Inr_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1041 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1042 |
lemma sats_Inr_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1043 |
"[| x \<in> nat; z \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1044 |
==> sats(A, Inr_fm(x,z), env) <-> is_Inr(**A, nth(x,env), nth(z,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1045 |
by (simp add: Inr_fm_def is_Inr_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1046 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1047 |
lemma Inr_iff_sats: |
13428 | 1048 |
"[| nth(i,env) = x; nth(k,env) = z; |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1049 |
i \<in> nat; k \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1050 |
==> is_Inr(**A, x, z) <-> sats(A, Inr_fm(i,k), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1051 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1052 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1053 |
theorem Inr_reflection: |
13428 | 1054 |
"REFLECTS[\<lambda>x. is_Inr(L,f(x),h(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1055 |
\<lambda>i x. is_Inr(**Lset(i),f(x),h(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1056 |
apply (simp only: is_Inr_def setclass_simps) |
13428 | 1057 |
apply (intro FOL_reflections function_reflections) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1058 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1059 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1060 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1061 |
subsubsection{*The Formula @{term is_Nil}, Internalized*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1062 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1063 |
(* is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1064 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1065 |
constdefs Nil_fm :: "i=>i" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1066 |
"Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))" |
13428 | 1067 |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1068 |
lemma Nil_type [TC]: "x \<in> nat ==> Nil_fm(x) \<in> formula" |
13428 | 1069 |
by (simp add: Nil_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1070 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1071 |
lemma sats_Nil_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1072 |
"[| x \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1073 |
==> sats(A, Nil_fm(x), env) <-> is_Nil(**A, nth(x,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1074 |
by (simp add: Nil_fm_def is_Nil_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1075 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1076 |
lemma Nil_iff_sats: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1077 |
"[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1078 |
==> is_Nil(**A, x) <-> sats(A, Nil_fm(i), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1079 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1080 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1081 |
theorem Nil_reflection: |
13428 | 1082 |
"REFLECTS[\<lambda>x. is_Nil(L,f(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1083 |
\<lambda>i x. is_Nil(**Lset(i),f(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1084 |
apply (simp only: is_Nil_def setclass_simps) |
13428 | 1085 |
apply (intro FOL_reflections function_reflections Inl_reflection) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1086 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1087 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1088 |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1089 |
subsubsection{*The Formula @{term is_Cons}, Internalized*} |
13395 | 1090 |
|
13387 | 1091 |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1092 |
(* "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1093 |
constdefs Cons_fm :: "[i,i,i]=>i" |
13428 | 1094 |
"Cons_fm(a,l,Z) == |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1095 |
Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1096 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1097 |
lemma Cons_type [TC]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1098 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> Cons_fm(x,y,z) \<in> formula" |
13428 | 1099 |
by (simp add: Cons_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1100 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1101 |
lemma sats_Cons_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1102 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
13428 | 1103 |
==> sats(A, Cons_fm(x,y,z), env) <-> |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1104 |
is_Cons(**A, nth(x,env), nth(y,env), nth(z,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1105 |
by (simp add: Cons_fm_def is_Cons_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1106 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1107 |
lemma Cons_iff_sats: |
13428 | 1108 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1109 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1110 |
==>is_Cons(**A, x, y, z) <-> sats(A, Cons_fm(i,j,k), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1111 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1112 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1113 |
theorem Cons_reflection: |
13428 | 1114 |
"REFLECTS[\<lambda>x. is_Cons(L,f(x),g(x),h(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1115 |
\<lambda>i x. is_Cons(**Lset(i),f(x),g(x),h(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1116 |
apply (simp only: is_Cons_def setclass_simps) |
13428 | 1117 |
apply (intro FOL_reflections pair_reflection Inr_reflection) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1118 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1119 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1120 |
subsubsection{*The Formula @{term is_quasilist}, Internalized*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1121 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1122 |
(* is_quasilist(M,xs) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1123 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1124 |
constdefs quasilist_fm :: "i=>i" |
13428 | 1125 |
"quasilist_fm(x) == |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1126 |
Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))" |
13428 | 1127 |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1128 |
lemma quasilist_type [TC]: "x \<in> nat ==> quasilist_fm(x) \<in> formula" |
13428 | 1129 |
by (simp add: quasilist_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1130 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1131 |
lemma sats_quasilist_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1132 |
"[| x \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1133 |
==> sats(A, quasilist_fm(x), env) <-> is_quasilist(**A, nth(x,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1134 |
by (simp add: quasilist_fm_def is_quasilist_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1135 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1136 |
lemma quasilist_iff_sats: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1137 |
"[| nth(i,env) = x; i \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1138 |
==> is_quasilist(**A, x) <-> sats(A, quasilist_fm(i), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1139 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1140 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1141 |
theorem quasilist_reflection: |
13428 | 1142 |
"REFLECTS[\<lambda>x. is_quasilist(L,f(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1143 |
\<lambda>i x. is_quasilist(**Lset(i),f(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1144 |
apply (simp only: is_quasilist_def setclass_simps) |
13428 | 1145 |
apply (intro FOL_reflections Nil_reflection Cons_reflection) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1146 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1147 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1148 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1149 |
subsection{*Absoluteness for the Function @{term nth}*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1150 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1151 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1152 |
subsubsection{*The Formula @{term is_tl}, Internalized*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1153 |
|
13428 | 1154 |
(* "is_tl(M,xs,T) == |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1155 |
(is_Nil(M,xs) --> T=xs) & |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1156 |
(\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) & |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1157 |
(is_quasilist(M,xs) | empty(M,T))" *) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1158 |
constdefs tl_fm :: "[i,i]=>i" |
13428 | 1159 |
"tl_fm(xs,T) == |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1160 |
And(Implies(Nil_fm(xs), Equal(T,xs)), |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1161 |
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))), |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1162 |
Or(quasilist_fm(xs), empty_fm(T))))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1163 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1164 |
lemma tl_type [TC]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1165 |
"[| x \<in> nat; y \<in> nat |] ==> tl_fm(x,y) \<in> formula" |
13428 | 1166 |
by (simp add: tl_fm_def) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1167 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1168 |
lemma sats_tl_fm [simp]: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1169 |
"[| x \<in> nat; y \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1170 |
==> sats(A, tl_fm(x,y), env) <-> is_tl(**A, nth(x,env), nth(y,env))" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1171 |
by (simp add: tl_fm_def is_tl_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1172 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1173 |
lemma tl_iff_sats: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1174 |
"[| nth(i,env) = x; nth(j,env) = y; |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1175 |
i \<in> nat; j \<in> nat; env \<in> list(A)|] |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1176 |
==> is_tl(**A, x, y) <-> sats(A, tl_fm(i,j), env)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1177 |
by simp |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1178 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1179 |
theorem tl_reflection: |
13428 | 1180 |
"REFLECTS[\<lambda>x. is_tl(L,f(x),g(x)), |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1181 |
\<lambda>i x. is_tl(**Lset(i),f(x),g(x))]" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1182 |
apply (simp only: is_tl_def setclass_simps) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1183 |
apply (intro FOL_reflections Nil_reflection Cons_reflection |
13428 | 1184 |
quasilist_reflection empty_reflection) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1185 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1186 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1187 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1188 |
subsubsection{*An Instance of Replacement for @{term nth}*} |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1189 |
|
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1190 |
lemma nth_replacement_Reflects: |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1191 |
"REFLECTS |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1192 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and> |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1193 |
is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)), |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1194 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
13428 | 1195 |
is_wfrec(**Lset(i), |
1196 |
iterates_MH(**Lset(i), |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1197 |
is_tl(**Lset(i)), z), memsn, u, y))]" |
13428 | 1198 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
1199 |
iterates_MH_reflection list_functor_reflection tl_reflection) |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1200 |
|
13428 | 1201 |
lemma nth_replacement: |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1202 |
"L(w) ==> iterates_replacement(L, %l t. is_tl(L,l,t), w)" |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1203 |
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) |
13428 | 1204 |
apply (rule strong_replacementI) |
1205 |
apply (rule rallI) |
|
1206 |
apply (rule separation_CollectI) |
|
1207 |
apply (subgoal_tac "L(Memrel(succ(n)))") |
|
1208 |
apply (rule_tac A="{A,n,w,z,Memrel(succ(n))}" in subset_LsetE, blast ) |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1209 |
apply (rule ReflectsE [OF nth_replacement_Reflects], assumption) |
13428 | 1210 |
apply (drule subset_Lset_ltD, assumption) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1211 |
apply (erule reflection_imp_L_separation) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1212 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
13428 | 1213 |
apply (elim conjE) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1214 |
apply (rule DPow_LsetI) |
13428 | 1215 |
apply (rename_tac v) |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1216 |
apply (rule bex_iff_sats conj_iff_sats)+ |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1217 |
apply (rule_tac env = "[u,v,A,z,w,Memrel(succ(n))]" in mem_iff_sats) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1218 |
apply (rule sep_rules | simp)+ |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1219 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1220 |
apply (rule sep_rules quasinat_iff_sats tl_iff_sats | simp)+ |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1221 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1222 |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1223 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1224 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1225 |
subsubsection{*Instantiating the locale @{text M_datatypes}*} |
13428 | 1226 |
|
1227 |
theorem M_datatypes_axioms_L: "M_datatypes_axioms(L)" |
|
1228 |
apply (rule M_datatypes_axioms.intro) |
|
1229 |
apply (assumption | rule |
|
1230 |
list_replacement1 list_replacement2 |
|
1231 |
formula_replacement1 formula_replacement2 |
|
1232 |
nth_replacement)+ |
|
1233 |
done |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1234 |
|
13428 | 1235 |
theorem M_datatypes_L: "PROP M_datatypes(L)" |
1236 |
apply (rule M_datatypes.intro) |
|
1237 |
apply (rule M_triv_axioms_L) |
|
1238 |
apply (rule M_axioms_axioms_L) |
|
1239 |
apply (rule M_trancl_axioms_L) |
|
1240 |
apply (rule M_wfrank_axioms_L) |
|
1241 |
apply (rule M_datatypes_axioms_L) |
|
1242 |
done |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1243 |
|
13428 | 1244 |
lemmas list_closed = M_datatypes.list_closed [OF M_datatypes_L] |
1245 |
and formula_closed = M_datatypes.formula_closed [OF M_datatypes_L] |
|
1246 |
and list_abs = M_datatypes.list_abs [OF M_datatypes_L] |
|
1247 |
and formula_abs = M_datatypes.formula_abs [OF M_datatypes_L] |
|
1248 |
and nth_abs = M_datatypes.nth_abs [OF M_datatypes_L] |
|
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1249 |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1250 |
declare list_closed [intro,simp] |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1251 |
declare formula_closed [intro,simp] |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1252 |
declare list_abs [simp] |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1253 |
declare formula_abs [simp] |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1254 |
declare nth_abs [simp] |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1255 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1256 |
|
13428 | 1257 |
subsection{*@{term L} is Closed Under the Operator @{term eclose}*} |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1258 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1259 |
subsubsection{*Instances of Replacement for @{term eclose}*} |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1260 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1261 |
lemma eclose_replacement1_Reflects: |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1262 |
"REFLECTS |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1263 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and> |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1264 |
is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)), |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1265 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
13428 | 1266 |
is_wfrec(**Lset(i), |
1267 |
iterates_MH(**Lset(i), big_union(**Lset(i)), A), |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1268 |
memsn, u, y))]" |
13428 | 1269 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
1270 |
iterates_MH_reflection) |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1271 |
|
13428 | 1272 |
lemma eclose_replacement1: |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1273 |
"L(A) ==> iterates_replacement(L, big_union(L), A)" |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1274 |
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) |
13428 | 1275 |
apply (rule strong_replacementI) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1276 |
apply (rule rallI) |
13428 | 1277 |
apply (rename_tac B) |
1278 |
apply (rule separation_CollectI) |
|
1279 |
apply (subgoal_tac "L(Memrel(succ(n)))") |
|
1280 |
apply (rule_tac A="{B,A,n,z,Memrel(succ(n))}" in subset_LsetE, blast ) |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1281 |
apply (rule ReflectsE [OF eclose_replacement1_Reflects], assumption) |
13428 | 1282 |
apply (drule subset_Lset_ltD, assumption) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1283 |
apply (erule reflection_imp_L_separation) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1284 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
13428 | 1285 |
apply (elim conjE) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1286 |
apply (rule DPow_LsetI) |
13428 | 1287 |
apply (rename_tac v) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1288 |
apply (rule bex_iff_sats conj_iff_sats)+ |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1289 |
apply (rule_tac env = "[u,v,A,n,B,Memrel(succ(n))]" in mem_iff_sats) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1290 |
apply (rule sep_rules | simp)+ |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1291 |
txt{*Can't get sat rules to work for higher-order operators, so just expand them!*} |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1292 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1293 |
apply (rule sep_rules big_union_iff_sats quasinat_iff_sats | simp)+ |
13409
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1294 |
done |
d4ea094c650e
Relativization and Separation for the function "nth"
paulson
parents:
13398
diff
changeset
|
1295 |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1296 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1297 |
lemma eclose_replacement2_Reflects: |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1298 |
"REFLECTS |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1299 |
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and> |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1300 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L, u, sn) \<and> membership(L, sn, msn) \<and> |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1301 |
is_wfrec (L, iterates_MH (L, big_union(L), A), |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1302 |
msn, u, x)), |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1303 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and> |
13428 | 1304 |
(\<exists>sn \<in> Lset(i). \<exists>msn \<in> Lset(i). |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1305 |
successor(**Lset(i), u, sn) \<and> membership(**Lset(i), sn, msn) \<and> |
13428 | 1306 |
is_wfrec (**Lset(i), |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1307 |
iterates_MH (**Lset(i), big_union(**Lset(i)), A), |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1308 |
msn, u, x))]" |
13428 | 1309 |
by (intro FOL_reflections function_reflections is_wfrec_reflection |
1310 |
iterates_MH_reflection) |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1311 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1312 |
|
13428 | 1313 |
lemma eclose_replacement2: |
1314 |
"L(A) ==> strong_replacement(L, |
|
1315 |
\<lambda>n y. n\<in>nat & |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1316 |
(\<exists>sn[L]. \<exists>msn[L]. successor(L,n,sn) & membership(L,sn,msn) & |
13428 | 1317 |
is_wfrec(L, iterates_MH(L,big_union(L), A), |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1318 |
msn, n, y)))" |
13428 | 1319 |
apply (rule strong_replacementI) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1320 |
apply (rule rallI) |
13428 | 1321 |
apply (rename_tac B) |
1322 |
apply (rule separation_CollectI) |
|
1323 |
apply (rule_tac A="{A,B,z,nat}" in subset_LsetE) |
|
1324 |
apply (blast intro: L_nat) |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1325 |
apply (rule ReflectsE [OF eclose_replacement2_Reflects], assumption) |
13428 | 1326 |
apply (drule subset_Lset_ltD, assumption) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1327 |
apply (erule reflection_imp_L_separation) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1328 |
apply (simp_all add: lt_Ord2) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1329 |
apply (rule DPow_LsetI) |
13428 | 1330 |
apply (rename_tac v) |
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1331 |
apply (rule bex_iff_sats conj_iff_sats)+ |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1332 |
apply (rule_tac env = "[u,v,A,B,nat]" in mem_iff_sats) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1333 |
apply (rule sep_rules | simp)+ |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1334 |
apply (simp add: is_wfrec_def M_is_recfun_def iterates_MH_def is_nat_case_def) |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1335 |
apply (rule sep_rules big_union_iff_sats quasinat_iff_sats | simp)+ |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1336 |
done |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1337 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1338 |
|
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1339 |
subsubsection{*Instantiating the locale @{text M_eclose}*} |
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1340 |
|
13428 | 1341 |
theorem M_eclose_axioms_L: "M_eclose_axioms(L)" |
1342 |
apply (rule M_eclose_axioms.intro) |
|
1343 |
apply (assumption | rule eclose_replacement1 eclose_replacement2)+ |
|
1344 |
done |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1345 |
|
13428 | 1346 |
theorem M_eclose_L: "PROP M_eclose(L)" |
1347 |
apply (rule M_eclose.intro) |
|
1348 |
apply (rule M_triv_axioms_L) |
|
1349 |
apply (rule M_axioms_axioms_L) |
|
1350 |
apply (rule M_trancl_axioms_L) |
|
1351 |
apply (rule M_wfrank_axioms_L) |
|
1352 |
apply (rule M_datatypes_axioms_L) |
|
1353 |
apply (rule M_eclose_axioms_L) |
|
1354 |
done |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1355 |
|
13428 | 1356 |
lemmas eclose_closed [intro, simp] = M_eclose.eclose_closed [OF M_eclose_L] |
1357 |
and eclose_abs [intro, simp] = M_eclose.eclose_abs [OF M_eclose_L] |
|
13422
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
paulson
parents:
13418
diff
changeset
|
1358 |
|
13348 | 1359 |
end |