author | paulson |
Thu, 15 Aug 2002 21:36:26 +0200 | |
changeset 13503 | d93f41fe35d2 |
parent 13502 | da43ebc02f17 |
child 13504 | 59066e97b551 |
permissions | -rw-r--r-- |
13494 | 1 |
(* Title: ZF/Constructible/Satisfies_absolute.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 2002 University of Cambridge |
|
5 |
*) |
|
6 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
7 |
header {*Absoluteness for the Satisfies Relation on Formulas*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
8 |
|
13494 | 9 |
theory Satisfies_absolute = Datatype_absolute + Rec_Separation: |
10 |
||
11 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
12 |
|
13494 | 13 |
subsubsection{*The Formula @{term is_depth}, Internalized*} |
14 |
||
15 |
(* "is_depth(M,p,n) == |
|
16 |
\<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M]. |
|
17 |
2 1 0 |
|
18 |
is_formula_N(M,n,formula_n) & p \<notin> formula_n & |
|
19 |
successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn" *) |
|
20 |
constdefs depth_fm :: "[i,i]=>i" |
|
21 |
"depth_fm(p,n) == |
|
22 |
Exists(Exists(Exists( |
|
23 |
And(formula_N_fm(n#+3,1), |
|
24 |
And(Neg(Member(p#+3,1)), |
|
25 |
And(succ_fm(n#+3,2), |
|
26 |
And(formula_N_fm(2,0), Member(p#+3,0))))))))" |
|
27 |
||
28 |
lemma depth_fm_type [TC]: |
|
29 |
"[| x \<in> nat; y \<in> nat |] ==> depth_fm(x,y) \<in> formula" |
|
30 |
by (simp add: depth_fm_def) |
|
31 |
||
32 |
lemma sats_depth_fm [simp]: |
|
33 |
"[| x \<in> nat; y < length(env); env \<in> list(A)|] |
|
34 |
==> sats(A, depth_fm(x,y), env) <-> |
|
35 |
is_depth(**A, nth(x,env), nth(y,env))" |
|
36 |
apply (frule_tac x=y in lt_length_in_nat, assumption) |
|
37 |
apply (simp add: depth_fm_def is_depth_def) |
|
38 |
done |
|
39 |
||
40 |
lemma depth_iff_sats: |
|
41 |
"[| nth(i,env) = x; nth(j,env) = y; |
|
42 |
i \<in> nat; j < length(env); env \<in> list(A)|] |
|
43 |
==> is_depth(**A, x, y) <-> sats(A, depth_fm(i,j), env)" |
|
44 |
by (simp add: sats_depth_fm) |
|
45 |
||
46 |
theorem depth_reflection: |
|
47 |
"REFLECTS[\<lambda>x. is_depth(L, f(x), g(x)), |
|
48 |
\<lambda>i x. is_depth(**Lset(i), f(x), g(x))]" |
|
49 |
apply (simp only: is_depth_def setclass_simps) |
|
50 |
apply (intro FOL_reflections function_reflections formula_N_reflection) |
|
51 |
done |
|
52 |
||
53 |
||
54 |
||
55 |
subsubsection{*The Operator @{term is_formula_case}*} |
|
56 |
||
57 |
text{*The arguments of @{term is_a} are always 2, 1, 0, and the formula |
|
58 |
will be enclosed by three quantifiers.*} |
|
59 |
||
60 |
(* is_formula_case :: |
|
61 |
"[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" |
|
62 |
"is_formula_case(M, is_a, is_b, is_c, is_d, v, z) == |
|
63 |
(\<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> is_Member(M,x,y,v) --> is_a(x,y,z)) & |
|
64 |
(\<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> is_Equal(M,x,y,v) --> is_b(x,y,z)) & |
|
65 |
(\<forall>x[M]. \<forall>y[M]. x\<in>formula --> y\<in>formula --> |
|
66 |
is_Nand(M,x,y,v) --> is_c(x,y,z)) & |
|
67 |
(\<forall>x[M]. x\<in>formula --> is_Forall(M,x,v) --> is_d(x,z))" *) |
|
68 |
||
69 |
constdefs formula_case_fm :: "[i, i, i, i, i, i]=>i" |
|
70 |
"formula_case_fm(is_a, is_b, is_c, is_d, v, z) == |
|
71 |
And(Forall(Forall(Implies(finite_ordinal_fm(1), |
|
72 |
Implies(finite_ordinal_fm(0), |
|
73 |
Implies(Member_fm(1,0,v#+2), |
|
74 |
Forall(Implies(Equal(0,z#+3), is_a))))))), |
|
75 |
And(Forall(Forall(Implies(finite_ordinal_fm(1), |
|
76 |
Implies(finite_ordinal_fm(0), |
|
77 |
Implies(Equal_fm(1,0,v#+2), |
|
78 |
Forall(Implies(Equal(0,z#+3), is_b))))))), |
|
79 |
And(Forall(Forall(Implies(mem_formula_fm(1), |
|
80 |
Implies(mem_formula_fm(0), |
|
81 |
Implies(Nand_fm(1,0,v#+2), |
|
82 |
Forall(Implies(Equal(0,z#+3), is_c))))))), |
|
83 |
Forall(Implies(mem_formula_fm(0), |
|
84 |
Implies(Forall_fm(0,succ(v)), |
|
85 |
Forall(Implies(Equal(0,z#+2), is_d))))))))" |
|
86 |
||
87 |
||
88 |
lemma is_formula_case_type [TC]: |
|
89 |
"[| is_a \<in> formula; is_b \<in> formula; is_c \<in> formula; is_d \<in> formula; |
|
90 |
x \<in> nat; y \<in> nat |] |
|
91 |
==> formula_case_fm(is_a, is_b, is_c, is_d, x, y) \<in> formula" |
|
92 |
by (simp add: formula_case_fm_def) |
|
93 |
||
94 |
lemma sats_formula_case_fm: |
|
95 |
assumes is_a_iff_sats: |
|
96 |
"!!a0 a1 a2. |
|
97 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
98 |
==> ISA(a2, a1, a0) <-> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
99 |
and is_b_iff_sats: |
|
100 |
"!!a0 a1 a2. |
|
101 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
102 |
==> ISB(a2, a1, a0) <-> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
103 |
and is_c_iff_sats: |
|
104 |
"!!a0 a1 a2. |
|
105 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
106 |
==> ISC(a2, a1, a0) <-> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
107 |
and is_d_iff_sats: |
|
108 |
"!!a0 a1. |
|
109 |
[|a0\<in>A; a1\<in>A|] |
|
110 |
==> ISD(a1, a0) <-> sats(A, is_d, Cons(a0,Cons(a1,env)))" |
|
111 |
shows |
|
112 |
"[|x \<in> nat; y < length(env); env \<in> list(A)|] |
|
113 |
==> sats(A, formula_case_fm(is_a,is_b,is_c,is_d,x,y), env) <-> |
|
114 |
is_formula_case(**A, ISA, ISB, ISC, ISD, nth(x,env), nth(y,env))" |
|
115 |
apply (frule_tac x=y in lt_length_in_nat, assumption) |
|
116 |
apply (simp add: formula_case_fm_def is_formula_case_def |
|
117 |
is_a_iff_sats [THEN iff_sym] is_b_iff_sats [THEN iff_sym] |
|
118 |
is_c_iff_sats [THEN iff_sym] is_d_iff_sats [THEN iff_sym]) |
|
119 |
done |
|
120 |
||
121 |
lemma formula_case_iff_sats: |
|
122 |
assumes is_a_iff_sats: |
|
123 |
"!!a0 a1 a2. |
|
124 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
125 |
==> ISA(a2, a1, a0) <-> sats(A, is_a, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
126 |
and is_b_iff_sats: |
|
127 |
"!!a0 a1 a2. |
|
128 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
129 |
==> ISB(a2, a1, a0) <-> sats(A, is_b, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
130 |
and is_c_iff_sats: |
|
131 |
"!!a0 a1 a2. |
|
132 |
[|a0\<in>A; a1\<in>A; a2\<in>A|] |
|
133 |
==> ISC(a2, a1, a0) <-> sats(A, is_c, Cons(a0,Cons(a1,Cons(a2,env))))" |
|
134 |
and is_d_iff_sats: |
|
135 |
"!!a0 a1. |
|
136 |
[|a0\<in>A; a1\<in>A|] |
|
137 |
==> ISD(a1, a0) <-> sats(A, is_d, Cons(a0,Cons(a1,env)))" |
|
138 |
shows |
|
139 |
"[|nth(i,env) = x; nth(j,env) = y; |
|
140 |
i \<in> nat; j < length(env); env \<in> list(A)|] |
|
141 |
==> is_formula_case(**A, ISA, ISB, ISC, ISD, x, y) <-> |
|
142 |
sats(A, formula_case_fm(is_a,is_b,is_c,is_d,i,j), env)" |
|
143 |
by (simp add: sats_formula_case_fm [OF is_a_iff_sats is_b_iff_sats |
|
144 |
is_c_iff_sats is_d_iff_sats]) |
|
145 |
||
146 |
||
147 |
text{*The second argument of @{term is_a} gives it direct access to @{term x}, |
|
148 |
which is essential for handling free variable references. Treatment is |
|
149 |
based on that of @{text is_nat_case_reflection}.*} |
|
150 |
theorem is_formula_case_reflection: |
|
151 |
assumes is_a_reflection: |
|
152 |
"!!h f g g'. REFLECTS[\<lambda>x. is_a(L, h(x), f(x), g(x), g'(x)), |
|
153 |
\<lambda>i x. is_a(**Lset(i), h(x), f(x), g(x), g'(x))]" |
|
154 |
and is_b_reflection: |
|
155 |
"!!h f g g'. REFLECTS[\<lambda>x. is_b(L, h(x), f(x), g(x), g'(x)), |
|
156 |
\<lambda>i x. is_b(**Lset(i), h(x), f(x), g(x), g'(x))]" |
|
157 |
and is_c_reflection: |
|
158 |
"!!h f g g'. REFLECTS[\<lambda>x. is_c(L, h(x), f(x), g(x), g'(x)), |
|
159 |
\<lambda>i x. is_c(**Lset(i), h(x), f(x), g(x), g'(x))]" |
|
160 |
and is_d_reflection: |
|
161 |
"!!h f g g'. REFLECTS[\<lambda>x. is_d(L, h(x), f(x), g(x)), |
|
162 |
\<lambda>i x. is_d(**Lset(i), h(x), f(x), g(x))]" |
|
163 |
shows "REFLECTS[\<lambda>x. is_formula_case(L, is_a(L,x), is_b(L,x), is_c(L,x), is_d(L,x), g(x), h(x)), |
|
164 |
\<lambda>i x. is_formula_case(**Lset(i), is_a(**Lset(i), x), is_b(**Lset(i), x), is_c(**Lset(i), x), is_d(**Lset(i), x), g(x), h(x))]" |
|
165 |
apply (simp (no_asm_use) only: is_formula_case_def setclass_simps) |
|
166 |
apply (intro FOL_reflections function_reflections finite_ordinal_reflection |
|
167 |
mem_formula_reflection |
|
168 |
Member_reflection Equal_reflection Nand_reflection Forall_reflection |
|
169 |
is_a_reflection is_b_reflection is_c_reflection is_d_reflection) |
|
170 |
done |
|
171 |
||
172 |
||
173 |
||
174 |
subsection {*Absoluteness for @{term formula_rec}*} |
|
175 |
||
176 |
constdefs |
|
177 |
||
178 |
formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" |
|
179 |
--{* the instance of @{term formula_case} in @{term formula_rec}*} |
|
180 |
"formula_rec_case(a,b,c,d,h) == |
|
181 |
formula_case (a, b, |
|
182 |
\<lambda>u v. c(u, v, h ` succ(depth(u)) ` u, |
|
183 |
h ` succ(depth(v)) ` v), |
|
184 |
\<lambda>u. d(u, h ` succ(depth(u)) ` u))" |
|
185 |
||
13502 | 186 |
is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" |
13503 | 187 |
--{* predicate to relativize the functional @{term formula_rec}*} |
13502 | 188 |
"is_formula_rec(M,MH,p,z) == |
13503 | 189 |
\<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & |
190 |
successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)" |
|
13494 | 191 |
|
192 |
text{*Unfold @{term formula_rec} to @{term formula_rec_case}*} |
|
193 |
lemma (in M_triv_axioms) formula_rec_eq2: |
|
194 |
"p \<in> formula ==> |
|
195 |
formula_rec(a,b,c,d,p) = |
|
196 |
transrec (succ(depth(p)), |
|
197 |
\<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h))) ` p" |
|
198 |
by (simp add: formula_rec_eq formula_rec_case_def) |
|
199 |
||
200 |
||
201 |
text{*Sufficient conditions to relative the instance of @{term formula_case} |
|
202 |
in @{term formula_rec}*} |
|
203 |
lemma (in M_datatypes) Relativize1_formula_rec_case: |
|
204 |
"[|Relativize2(M, nat, nat, is_a, a); |
|
205 |
Relativize2(M, nat, nat, is_b, b); |
|
206 |
Relativize2 (M, formula, formula, |
|
207 |
is_c, \<lambda>u v. c(u, v, h`succ(depth(u))`u, h`succ(depth(v))`v)); |
|
208 |
Relativize1(M, formula, |
|
209 |
is_d, \<lambda>u. d(u, h ` succ(depth(u)) ` u)); |
|
210 |
M(h) |] |
|
211 |
==> Relativize1(M, formula, |
|
212 |
is_formula_case (M, is_a, is_b, is_c, is_d), |
|
213 |
formula_rec_case(a, b, c, d, h))" |
|
214 |
apply (simp (no_asm) add: formula_rec_case_def Relativize1_def) |
|
215 |
apply (simp add: formula_case_abs) |
|
216 |
done |
|
217 |
||
218 |
||
219 |
text{*This locale packages the premises of the following theorems, |
|
220 |
which is the normal purpose of locales. It doesn't accumulate |
|
221 |
constraints on the class @{term M}, as in most of this deveopment.*} |
|
222 |
locale M_formula_rec = M_eclose + |
|
223 |
fixes a and is_a and b and is_b and c and is_c and d and is_d and MH |
|
224 |
defines |
|
225 |
"MH(u::i,f,z) == |
|
13502 | 226 |
\<forall>fml[M]. is_formula(M,fml) --> |
227 |
is_lambda |
|
228 |
(M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)" |
|
13494 | 229 |
|
230 |
assumes a_closed: "[|x\<in>nat; y\<in>nat|] ==> M(a(x,y))" |
|
231 |
and a_rel: "Relativize2(M, nat, nat, is_a, a)" |
|
232 |
and b_closed: "[|x\<in>nat; y\<in>nat|] ==> M(b(x,y))" |
|
233 |
and b_rel: "Relativize2(M, nat, nat, is_b, b)" |
|
234 |
and c_closed: "[|x \<in> formula; y \<in> formula; M(gx); M(gy)|] |
|
235 |
==> M(c(x, y, gx, gy))" |
|
236 |
and c_rel: |
|
237 |
"M(f) ==> |
|
238 |
Relativize2 (M, formula, formula, is_c(f), |
|
239 |
\<lambda>u v. c(u, v, f ` succ(depth(u)) ` u, f ` succ(depth(v)) ` v))" |
|
240 |
and d_closed: "[|x \<in> formula; M(gx)|] ==> M(d(x, gx))" |
|
241 |
and d_rel: |
|
242 |
"M(f) ==> |
|
243 |
Relativize1(M, formula, is_d(f), \<lambda>u. d(u, f ` succ(depth(u)) ` u))" |
|
244 |
and fr_replace: "n \<in> nat ==> transrec_replacement(M,MH,n)" |
|
245 |
and fr_lam_replace: |
|
246 |
"M(g) ==> |
|
247 |
strong_replacement |
|
248 |
(M, \<lambda>x y. x \<in> formula & |
|
249 |
y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)"; |
|
250 |
||
251 |
lemma (in M_formula_rec) formula_rec_case_closed: |
|
252 |
"[|M(g); p \<in> formula|] ==> M(formula_rec_case(a, b, c, d, g, p))" |
|
253 |
by (simp add: formula_rec_case_def a_closed b_closed c_closed d_closed) |
|
254 |
||
255 |
lemma (in M_formula_rec) formula_rec_lam_closed: |
|
256 |
"M(g) ==> M(Lambda (formula, formula_rec_case(a,b,c,d,g)))" |
|
257 |
by (simp add: lam_closed2 fr_lam_replace formula_rec_case_closed) |
|
258 |
||
259 |
lemma (in M_formula_rec) MH_rel2: |
|
260 |
"relativize2 (M, MH, |
|
261 |
\<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h)))" |
|
262 |
apply (simp add: relativize2_def MH_def, clarify) |
|
263 |
apply (rule lambda_abs2) |
|
264 |
apply (rule fr_lam_replace, assumption) |
|
265 |
apply (rule Relativize1_formula_rec_case) |
|
266 |
apply (simp_all add: a_rel b_rel c_rel d_rel formula_rec_case_closed) |
|
267 |
done |
|
268 |
||
269 |
lemma (in M_formula_rec) fr_transrec_closed: |
|
270 |
"n \<in> nat |
|
271 |
==> M(transrec |
|
272 |
(n, \<lambda>x h. Lambda(formula, formula_rec_case(a, b, c, d, h))))" |
|
273 |
by (simp add: transrec_closed [OF fr_replace MH_rel2] |
|
274 |
nat_into_M formula_rec_lam_closed) |
|
275 |
||
276 |
text{*The main two results: @{term formula_rec} is absolute for @{term M}.*} |
|
277 |
theorem (in M_formula_rec) formula_rec_closed: |
|
278 |
"p \<in> formula ==> M(formula_rec(a,b,c,d,p))" |
|
279 |
by (simp add: formula_rec_eq2 fr_transrec_closed |
|
280 |
transM [OF _ formula_closed]) |
|
281 |
||
282 |
theorem (in M_formula_rec) formula_rec_abs: |
|
283 |
"[| p \<in> formula; M(z)|] |
|
13502 | 284 |
==> is_formula_rec(M,MH,p,z) <-> z = formula_rec(a,b,c,d,p)" |
13494 | 285 |
by (simp add: is_formula_rec_def formula_rec_eq2 transM [OF _ formula_closed] |
13503 | 286 |
transrec_abs [OF fr_replace MH_rel2] depth_type |
13494 | 287 |
fr_transrec_closed formula_rec_lam_closed eq_commute) |
288 |
||
289 |
||
290 |
subsection {*Absoluteness for the Function @{term satisfies}*} |
|
291 |
||
292 |
constdefs |
|
293 |
is_depth_apply :: "[i=>o,i,i,i] => o" |
|
294 |
--{*Merely a useful abbreviation for the sequel.*} |
|
295 |
"is_depth_apply(M,h,p,z) == |
|
296 |
\<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
297 |
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) & |
13494 | 298 |
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)" |
299 |
||
300 |
lemma (in M_datatypes) is_depth_apply_abs [simp]: |
|
301 |
"[|M(h); p \<in> formula; M(z)|] |
|
302 |
==> is_depth_apply(M,h,p,z) <-> z = h ` succ(depth(p)) ` p" |
|
303 |
by (simp add: is_depth_apply_def formula_into_M depth_type eq_commute) |
|
304 |
||
305 |
||
306 |
||
307 |
text{*There is at present some redundancy between the relativizations in |
|
308 |
e.g. @{text satisfies_is_a} and those in e.g. @{text Member_replacement}.*} |
|
309 |
||
310 |
text{*These constants let us instantiate the parameters @{term a}, @{term b}, |
|
311 |
@{term c}, @{term d}, etc., of the locale @{text M_formula_rec}.*} |
|
312 |
constdefs |
|
313 |
satisfies_a :: "[i,i,i]=>i" |
|
314 |
"satisfies_a(A) == |
|
315 |
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env))" |
|
316 |
||
317 |
satisfies_is_a :: "[i=>o,i,i,i,i]=>o" |
|
318 |
"satisfies_is_a(M,A) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
319 |
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
320 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
321 |
\<lambda>env z. is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
322 |
\<exists>nx[M]. \<exists>ny[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
323 |
is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
324 |
zz)" |
13494 | 325 |
|
326 |
satisfies_b :: "[i,i,i]=>i" |
|
327 |
"satisfies_b(A) == |
|
328 |
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env))" |
|
329 |
||
330 |
satisfies_is_b :: "[i=>o,i,i,i,i]=>o" |
|
331 |
--{*We simplify the formula to have just @{term nx} rather than |
|
332 |
introducing @{term ny} with @{term "nx=ny"} *} |
|
333 |
"satisfies_is_b(M,A) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
334 |
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
335 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
336 |
\<lambda>env z. is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
337 |
\<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
338 |
zz)" |
13494 | 339 |
|
340 |
satisfies_c :: "[i,i,i,i,i]=>i" |
|
13502 | 341 |
"satisfies_c(A) == \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)" |
13494 | 342 |
|
343 |
satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" |
|
344 |
"satisfies_is_c(M,A,h) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
345 |
\<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
346 |
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. |
13494 | 347 |
(\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & |
348 |
(\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
349 |
(\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
350 |
zz)" |
13494 | 351 |
|
352 |
satisfies_d :: "[i,i,i]=>i" |
|
353 |
"satisfies_d(A) |
|
354 |
== \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)" |
|
355 |
||
356 |
satisfies_is_d :: "[i=>o,i,i,i,i]=>o" |
|
357 |
"satisfies_is_d(M,A,h) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
358 |
\<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
359 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
360 |
\<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
361 |
is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
362 |
\<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
363 |
x\<in>A --> is_Cons(M,x,env,xenv) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
364 |
fun_apply(M,rp,xenv,hp) --> number1(M,hp), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
365 |
z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
366 |
zz)" |
13494 | 367 |
|
368 |
satisfies_MH :: "[i=>o,i,i,i,i]=>o" |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
369 |
--{*The variable @{term u} is unused, but gives @{term satisfies_MH} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
370 |
the correct arity.*} |
13494 | 371 |
"satisfies_MH == |
13502 | 372 |
\<lambda>M A u f z. |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
373 |
\<forall>fml[M]. is_formula(M,fml) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
374 |
is_lambda (M, fml, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
375 |
is_formula_case (M, satisfies_is_a(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
376 |
satisfies_is_b(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
377 |
satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)), |
13502 | 378 |
z)" |
13494 | 379 |
|
13502 | 380 |
is_satisfies :: "[i=>o,i,i,i]=>o" |
13503 | 381 |
"is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))" |
13502 | 382 |
|
383 |
||
384 |
text{*This lemma relates the fragments defined above to the original primitive |
|
385 |
recursion in @{term satisfies}. |
|
386 |
Induction is not required: the definitions are directly equal!*} |
|
387 |
lemma satisfies_eq: |
|
388 |
"satisfies(A,p) = |
|
389 |
formula_rec (satisfies_a(A), satisfies_b(A), |
|
390 |
satisfies_c(A), satisfies_d(A), p)" |
|
391 |
by (simp add: satisfies_formula_def satisfies_a_def satisfies_b_def |
|
392 |
satisfies_c_def satisfies_d_def) |
|
13494 | 393 |
|
394 |
text{*Further constraints on the class @{term M} in order to prove |
|
395 |
absoluteness for the constants defined above. The ultimate goal |
|
396 |
is the absoluteness of the function @{term satisfies}. *} |
|
13502 | 397 |
locale M_satisfies = M_eclose + |
13494 | 398 |
assumes |
399 |
Member_replacement: |
|
400 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
401 |
==> strong_replacement |
|
402 |
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. |
|
403 |
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & |
|
404 |
is_bool_of_o(M, nx \<in> ny, bo) & |
|
405 |
pair(M, env, bo, z))" |
|
406 |
and |
|
407 |
Equal_replacement: |
|
408 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
409 |
==> strong_replacement |
|
410 |
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. |
|
411 |
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & |
|
412 |
is_bool_of_o(M, nx = ny, bo) & |
|
413 |
pair(M, env, bo, z))" |
|
414 |
and |
|
415 |
Nand_replacement: |
|
416 |
"[|M(A); M(rp); M(rq)|] |
|
417 |
==> strong_replacement |
|
418 |
(M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M]. |
|
419 |
fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) & |
|
420 |
is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) & |
|
421 |
env \<in> list(A) & pair(M, env, notpq, z))" |
|
422 |
and |
|
423 |
Forall_replacement: |
|
424 |
"[|M(A); M(rp)|] |
|
425 |
==> strong_replacement |
|
426 |
(M, \<lambda>env z. \<exists>bo[M]. |
|
427 |
env \<in> list(A) & |
|
428 |
is_bool_of_o (M, |
|
429 |
\<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. |
|
430 |
a\<in>A --> is_Cons(M,a,env,co) --> |
|
431 |
fun_apply(M,rp,co,rpco) --> number1(M, rpco), |
|
432 |
bo) & |
|
433 |
pair(M,env,bo,z))" |
|
434 |
and |
|
435 |
formula_rec_replacement: |
|
436 |
--{*For the @{term transrec}*} |
|
437 |
"[|n \<in> nat; M(A)|] ==> transrec_replacement(M, satisfies_MH(M,A), n)" |
|
438 |
and |
|
439 |
formula_rec_lambda_replacement: |
|
440 |
--{*For the @{text "\<lambda>-abstraction"} in the @{term transrec} body*} |
|
13502 | 441 |
"[|M(g); M(A)|] ==> |
442 |
strong_replacement (M, |
|
443 |
\<lambda>x y. mem_formula(M,x) & |
|
444 |
(\<exists>c[M]. is_formula_case(M, satisfies_is_a(M,A), |
|
445 |
satisfies_is_b(M,A), |
|
446 |
satisfies_is_c(M,A,g), |
|
447 |
satisfies_is_d(M,A,g), x, c) & |
|
448 |
pair(M, x, c, y)))" |
|
13494 | 449 |
|
450 |
||
451 |
lemma (in M_satisfies) Member_replacement': |
|
452 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
453 |
==> strong_replacement |
|
454 |
(M, \<lambda>env z. env \<in> list(A) & |
|
455 |
z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)" |
|
456 |
by (insert Member_replacement, simp) |
|
457 |
||
458 |
lemma (in M_satisfies) Equal_replacement': |
|
459 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
460 |
==> strong_replacement |
|
461 |
(M, \<lambda>env z. env \<in> list(A) & |
|
462 |
z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)" |
|
463 |
by (insert Equal_replacement, simp) |
|
464 |
||
465 |
lemma (in M_satisfies) Nand_replacement': |
|
466 |
"[|M(A); M(rp); M(rq)|] |
|
467 |
==> strong_replacement |
|
468 |
(M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)" |
|
469 |
by (insert Nand_replacement, simp) |
|
470 |
||
471 |
lemma (in M_satisfies) Forall_replacement': |
|
472 |
"[|M(A); M(rp)|] |
|
473 |
==> strong_replacement |
|
474 |
(M, \<lambda>env z. |
|
475 |
env \<in> list(A) & |
|
476 |
z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)" |
|
477 |
by (insert Forall_replacement, simp) |
|
478 |
||
479 |
lemma (in M_satisfies) a_closed: |
|
480 |
"[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_a(A,x,y))" |
|
481 |
apply (simp add: satisfies_a_def) |
|
482 |
apply (blast intro: lam_closed2 Member_replacement') |
|
483 |
done |
|
484 |
||
485 |
lemma (in M_satisfies) a_rel: |
|
486 |
"M(A) ==> Relativize2(M, nat, nat, satisfies_is_a(M,A), satisfies_a(A))" |
|
487 |
apply (simp add: Relativize2_def satisfies_is_a_def satisfies_a_def) |
|
488 |
apply (simp add: lambda_abs2 [OF Member_replacement'] Relativize1_def) |
|
489 |
done |
|
490 |
||
491 |
lemma (in M_satisfies) b_closed: |
|
492 |
"[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_b(A,x,y))" |
|
493 |
apply (simp add: satisfies_b_def) |
|
494 |
apply (blast intro: lam_closed2 Equal_replacement') |
|
495 |
done |
|
496 |
||
497 |
lemma (in M_satisfies) b_rel: |
|
498 |
"M(A) ==> Relativize2(M, nat, nat, satisfies_is_b(M,A), satisfies_b(A))" |
|
499 |
apply (simp add: Relativize2_def satisfies_is_b_def satisfies_b_def) |
|
500 |
apply (simp add: lambda_abs2 [OF Equal_replacement'] Relativize1_def) |
|
501 |
done |
|
502 |
||
503 |
lemma (in M_satisfies) c_closed: |
|
504 |
"[|M(A); x \<in> formula; y \<in> formula; M(rx); M(ry)|] |
|
505 |
==> M(satisfies_c(A,x,y,rx,ry))" |
|
506 |
apply (simp add: satisfies_c_def) |
|
507 |
apply (rule lam_closed2) |
|
508 |
apply (rule Nand_replacement') |
|
509 |
apply (simp_all add: formula_into_M list_into_M [of _ A]) |
|
510 |
done |
|
511 |
||
512 |
lemma (in M_satisfies) c_rel: |
|
513 |
"[|M(A); M(f)|] ==> |
|
514 |
Relativize2 (M, formula, formula, |
|
515 |
satisfies_is_c(M,A,f), |
|
516 |
\<lambda>u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u, |
|
517 |
f ` succ(depth(v)) ` v))" |
|
518 |
apply (simp add: Relativize2_def satisfies_is_c_def satisfies_c_def) |
|
519 |
apply (simp add: lambda_abs2 [OF Nand_replacement' triv_Relativize1] |
|
520 |
formula_into_M) |
|
521 |
done |
|
522 |
||
523 |
lemma (in M_satisfies) d_closed: |
|
524 |
"[|M(A); x \<in> formula; M(rx)|] ==> M(satisfies_d(A,x,rx))" |
|
525 |
apply (simp add: satisfies_d_def) |
|
526 |
apply (rule lam_closed2) |
|
527 |
apply (rule Forall_replacement') |
|
528 |
apply (simp_all add: formula_into_M list_into_M [of _ A]) |
|
529 |
done |
|
530 |
||
531 |
lemma (in M_satisfies) d_rel: |
|
532 |
"[|M(A); M(f)|] ==> |
|
533 |
Relativize1(M, formula, satisfies_is_d(M,A,f), |
|
534 |
\<lambda>u. satisfies_d(A, u, f ` succ(depth(u)) ` u))" |
|
535 |
apply (simp del: rall_abs |
|
536 |
add: Relativize1_def satisfies_is_d_def satisfies_d_def) |
|
537 |
apply (simp add: lambda_abs2 [OF Forall_replacement' triv_Relativize1] |
|
538 |
formula_into_M) |
|
539 |
done |
|
540 |
||
541 |
||
542 |
lemma (in M_satisfies) fr_replace: |
|
543 |
"[|n \<in> nat; M(A)|] ==> transrec_replacement(M,satisfies_MH(M,A),n)" |
|
544 |
by (blast intro: formula_rec_replacement) |
|
545 |
||
13502 | 546 |
lemma (in M_satisfies) formula_case_satisfies_closed: |
547 |
"[|M(g); M(A); x \<in> formula|] ==> |
|
548 |
M(formula_case (satisfies_a(A), satisfies_b(A), |
|
549 |
\<lambda>u v. satisfies_c(A, u, v, |
|
550 |
g ` succ(depth(u)) ` u, g ` succ(depth(v)) ` v), |
|
551 |
\<lambda>u. satisfies_d (A, u, g ` succ(depth(u)) ` u), |
|
552 |
x))" |
|
553 |
by (blast intro: formula_case_closed a_closed b_closed c_closed d_closed) |
|
554 |
||
13494 | 555 |
lemma (in M_satisfies) fr_lam_replace: |
13502 | 556 |
"[|M(g); M(A)|] ==> |
13494 | 557 |
strong_replacement (M, \<lambda>x y. x \<in> formula & |
558 |
y = \<langle>x, |
|
559 |
formula_rec_case(satisfies_a(A), |
|
560 |
satisfies_b(A), |
|
561 |
satisfies_c(A), |
|
562 |
satisfies_d(A), g, x)\<rangle>)" |
|
13502 | 563 |
apply (insert formula_rec_lambda_replacement) |
564 |
apply (simp add: formula_rec_case_def formula_case_satisfies_closed |
|
565 |
formula_case_abs [OF a_rel b_rel c_rel d_rel]) |
|
566 |
done |
|
13494 | 567 |
|
568 |
||
569 |
||
13502 | 570 |
text{*Instantiate locale @{text M_formula_rec} for the |
571 |
Function @{term satisfies}*} |
|
13494 | 572 |
|
13502 | 573 |
lemma (in M_satisfies) M_formula_rec_axioms_M: |
574 |
"M(A) ==> |
|
575 |
M_formula_rec_axioms(M, satisfies_a(A), satisfies_is_a(M,A), |
|
576 |
satisfies_b(A), satisfies_is_b(M,A), |
|
577 |
satisfies_c(A), satisfies_is_c(M,A), |
|
578 |
satisfies_d(A), satisfies_is_d(M,A))" |
|
579 |
apply (rule M_formula_rec_axioms.intro) |
|
580 |
apply (assumption | |
|
581 |
rule a_closed a_rel b_closed b_rel c_closed c_rel d_closed d_rel |
|
582 |
fr_replace [unfolded satisfies_MH_def] |
|
583 |
fr_lam_replace) + |
|
13494 | 584 |
done |
585 |
||
586 |
||
13502 | 587 |
theorem (in M_satisfies) M_formula_rec_M: |
588 |
"M(A) ==> |
|
589 |
PROP M_formula_rec(M, satisfies_a(A), satisfies_is_a(M,A), |
|
590 |
satisfies_b(A), satisfies_is_b(M,A), |
|
591 |
satisfies_c(A), satisfies_is_c(M,A), |
|
592 |
satisfies_d(A), satisfies_is_d(M,A))" |
|
593 |
apply (rule M_formula_rec.intro, assumption+) |
|
594 |
apply (erule M_formula_rec_axioms_M) |
|
13494 | 595 |
done |
596 |
||
13502 | 597 |
lemmas (in M_satisfies) |
598 |
satisfies_closed = M_formula_rec.formula_rec_closed [OF M_formula_rec_M] |
|
599 |
and |
|
600 |
satisfies_abs = M_formula_rec.formula_rec_abs [OF M_formula_rec_M]; |
|
13494 | 601 |
|
602 |
||
13502 | 603 |
lemma (in M_satisfies) satisfies_closed: |
604 |
"[|M(A); p \<in> formula|] ==> M(satisfies(A,p))" |
|
605 |
by (simp add: M_formula_rec.formula_rec_closed [OF M_formula_rec_M] |
|
606 |
satisfies_eq) |
|
13494 | 607 |
|
13502 | 608 |
lemma (in M_satisfies) satisfies_abs: |
609 |
"[|M(A); M(z); p \<in> formula|] |
|
610 |
==> is_satisfies(M,A,p,z) <-> z = satisfies(A,p)" |
|
611 |
by (simp only: M_formula_rec.formula_rec_abs [OF M_formula_rec_M] |
|
13503 | 612 |
satisfies_eq is_satisfies_def satisfies_MH_def) |
13494 | 613 |
|
614 |
||
13502 | 615 |
subsection{*Internalizations Needed to Instantiate @{text "M_satisfies"}*} |
13494 | 616 |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
617 |
subsubsection{*The Operator @{term is_depth_apply}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
618 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
619 |
(* is_depth_apply(M,h,p,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
620 |
\<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
621 |
2 1 0 |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
622 |
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) & |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
623 |
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
624 |
constdefs depth_apply_fm :: "[i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
625 |
"depth_apply_fm(h,p,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
626 |
Exists(Exists(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
627 |
And(finite_ordinal_fm(2), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
628 |
And(depth_fm(p#+3,2), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
629 |
And(succ_fm(2,1), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
630 |
And(fun_apply_fm(h#+3,1,0), fun_apply_fm(0,p#+3,z#+3))))))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
631 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
632 |
lemma depth_apply_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
633 |
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> depth_apply_fm(x,y,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
634 |
by (simp add: depth_apply_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
635 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
636 |
lemma sats_depth_apply_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
637 |
"[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
638 |
==> sats(A, depth_apply_fm(x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
639 |
is_depth_apply(**A, nth(x,env), nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
640 |
by (simp add: depth_apply_fm_def is_depth_apply_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
641 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
642 |
lemma depth_apply_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
643 |
"[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
644 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
645 |
==> is_depth_apply(**A, x, y, z) <-> sats(A, depth_apply_fm(i,j,k), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
646 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
647 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
648 |
lemma depth_apply_reflection: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
649 |
"REFLECTS[\<lambda>x. is_depth_apply(L,f(x),g(x),h(x)), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
650 |
\<lambda>i x. is_depth_apply(**Lset(i),f(x),g(x),h(x))]" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
651 |
apply (simp only: is_depth_apply_def setclass_simps) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
652 |
apply (intro FOL_reflections function_reflections depth_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
653 |
finite_ordinal_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
654 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
655 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
656 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
657 |
subsubsection{*The Operator @{term satisfies_is_a}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
658 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
659 |
(* satisfies_is_a(M,A) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
660 |
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
661 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
662 |
\<lambda>env z. is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
663 |
\<exists>nx[M]. \<exists>ny[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
664 |
is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
665 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
666 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
667 |
constdefs satisfies_is_a_fm :: "[i,i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
668 |
"satisfies_is_a_fm(A,x,y,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
669 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
670 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
671 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
672 |
bool_of_o_fm(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
673 |
Exists(And(nth_fm(x#+6,3,1), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
674 |
And(nth_fm(y#+6,3,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
675 |
Member(1,0))))), 0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
676 |
0, succ(z))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
677 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
678 |
lemma satisfies_is_a_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
679 |
"[| A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
680 |
==> satisfies_is_a_fm(A,x,y,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
681 |
by (simp add: satisfies_is_a_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
682 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
683 |
lemma sats_satisfies_is_a_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
684 |
"[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
685 |
==> sats(A, satisfies_is_a_fm(u,x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
686 |
satisfies_is_a(**A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
687 |
apply (frule_tac x=x in lt_length_in_nat, assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
688 |
apply (frule_tac x=y in lt_length_in_nat, assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
689 |
apply (simp add: satisfies_is_a_fm_def satisfies_is_a_def sats_lambda_fm |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
690 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
691 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
692 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
693 |
lemma satisfies_is_a_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
694 |
"[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
695 |
u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
696 |
==> satisfies_is_a(**A,nu,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
697 |
sats(A, satisfies_is_a_fm(u,x,y,z), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
698 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
699 |
|
13494 | 700 |
theorem satisfies_is_a_reflection: |
701 |
"REFLECTS[\<lambda>x. satisfies_is_a(L,f(x),g(x),h(x),g'(x)), |
|
702 |
\<lambda>i x. satisfies_is_a(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
703 |
apply (unfold satisfies_is_a_def) |
|
704 |
apply (intro FOL_reflections is_lambda_reflection bool_of_o_reflection |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
705 |
nth_reflection is_list_reflection) |
13494 | 706 |
done |
707 |
||
708 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
709 |
subsubsection{*The Operator @{term satisfies_is_b}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
710 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
711 |
(* satisfies_is_b(M,A) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
712 |
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
713 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
714 |
\<lambda>env z. is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
715 |
\<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
716 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
717 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
718 |
constdefs satisfies_is_b_fm :: "[i,i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
719 |
"satisfies_is_b_fm(A,x,y,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
720 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
721 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
722 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
723 |
bool_of_o_fm(Exists(And(nth_fm(x#+5,2,0), nth_fm(y#+5,2,0))), 0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
724 |
0, succ(z))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
725 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
726 |
lemma satisfies_is_b_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
727 |
"[| A \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
728 |
==> satisfies_is_b_fm(A,x,y,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
729 |
by (simp add: satisfies_is_b_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
730 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
731 |
lemma sats_satisfies_is_b_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
732 |
"[| u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
733 |
==> sats(A, satisfies_is_b_fm(u,x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
734 |
satisfies_is_b(**A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
735 |
apply (frule_tac x=x in lt_length_in_nat, assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
736 |
apply (frule_tac x=y in lt_length_in_nat, assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
737 |
apply (simp add: satisfies_is_b_fm_def satisfies_is_b_def sats_lambda_fm |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
738 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
739 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
740 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
741 |
lemma satisfies_is_b_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
742 |
"[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
743 |
u \<in> nat; x < length(env); y < length(env); z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
744 |
==> satisfies_is_b(**A,nu,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
745 |
sats(A, satisfies_is_b_fm(u,x,y,z), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
746 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
747 |
|
13494 | 748 |
theorem satisfies_is_b_reflection: |
749 |
"REFLECTS[\<lambda>x. satisfies_is_b(L,f(x),g(x),h(x),g'(x)), |
|
750 |
\<lambda>i x. satisfies_is_b(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
751 |
apply (unfold satisfies_is_b_def) |
|
752 |
apply (intro FOL_reflections is_lambda_reflection bool_of_o_reflection |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
753 |
nth_reflection is_list_reflection) |
13494 | 754 |
done |
755 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
756 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
757 |
subsubsection{*The Operator @{term satisfies_is_c}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
758 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
759 |
(* satisfies_is_c(M,A,h) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
760 |
\<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
761 |
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
762 |
(\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
763 |
(\<exists>rq[M]. is_depth_apply(M,h,q,rq) & fun_apply(M,rq,env,hq)) & |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
764 |
(\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
765 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
766 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
767 |
constdefs satisfies_is_c_fm :: "[i,i,i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
768 |
"satisfies_is_c_fm(A,h,p,q,zz) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
769 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
770 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
771 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
772 |
Exists(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
773 |
And(Exists(And(depth_apply_fm(h#+7,p#+7,0), fun_apply_fm(0,4,2))), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
774 |
And(Exists(And(depth_apply_fm(h#+7,q#+7,0), fun_apply_fm(0,4,1))), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
775 |
Exists(And(and_fm(2,1,0), not_fm(0,3))))))), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
776 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
777 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
778 |
lemma satisfies_is_c_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
779 |
"[| A \<in> nat; h \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat |] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
780 |
==> satisfies_is_c_fm(A,h,x,y,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
781 |
by (simp add: satisfies_is_c_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
782 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
783 |
lemma sats_satisfies_is_c_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
784 |
"[| u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
785 |
==> sats(A, satisfies_is_c_fm(u,v,x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
786 |
satisfies_is_c(**A, nth(u,env), nth(v,env), nth(x,env), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
787 |
nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
788 |
by (simp add: satisfies_is_c_fm_def satisfies_is_c_def sats_lambda_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
789 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
790 |
lemma satisfies_is_c_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
791 |
"[| nth(u,env) = nu; nth(v,env) = nv; nth(x,env) = nx; nth(y,env) = ny; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
792 |
nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
793 |
u \<in> nat; v \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
794 |
==> satisfies_is_c(**A,nu,nv,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
795 |
sats(A, satisfies_is_c_fm(u,v,x,y,z), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
796 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
797 |
|
13494 | 798 |
theorem satisfies_is_c_reflection: |
799 |
"REFLECTS[\<lambda>x. satisfies_is_c(L,f(x),g(x),h(x),g'(x),h'(x)), |
|
800 |
\<lambda>i x. satisfies_is_c(**Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
801 |
apply (unfold satisfies_is_c_def) |
13494 | 802 |
apply (intro FOL_reflections function_reflections is_lambda_reflection |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
803 |
extra_reflections nth_reflection depth_apply_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
804 |
is_list_reflection) |
13494 | 805 |
done |
806 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
807 |
subsubsection{*The Operator @{term satisfies_is_d}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
808 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
809 |
(* satisfies_is_d(M,A,h) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
810 |
\<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
811 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
812 |
\<lambda>env z. \<exists>rp[M]. is_depth_apply(M,h,p,rp) & |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
813 |
is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
814 |
\<forall>x[M]. \<forall>xenv[M]. \<forall>hp[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
815 |
x\<in>A --> is_Cons(M,x,env,xenv) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
816 |
fun_apply(M,rp,xenv,hp) --> number1(M,hp), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
817 |
z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
818 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
819 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
820 |
constdefs satisfies_is_d_fm :: "[i,i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
821 |
"satisfies_is_d_fm(A,h,p,zz) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
822 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
823 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
824 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
825 |
Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
826 |
And(depth_apply_fm(h#+5,p#+5,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
827 |
bool_of_o_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
828 |
Forall(Forall(Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
829 |
Implies(Member(2,A#+8), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
830 |
Implies(Cons_fm(2,5,1), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
831 |
Implies(fun_apply_fm(3,1,0), number1_fm(0))))))), 1))), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
832 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
833 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
834 |
lemma satisfies_is_d_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
835 |
"[| A \<in> nat; h \<in> nat; x \<in> nat; z \<in> nat |] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
836 |
==> satisfies_is_d_fm(A,h,x,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
837 |
by (simp add: satisfies_is_d_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
838 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
839 |
lemma sats_satisfies_is_d_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
840 |
"[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
841 |
==> sats(A, satisfies_is_d_fm(u,x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
842 |
satisfies_is_d(**A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
843 |
by (simp add: satisfies_is_d_fm_def satisfies_is_d_def sats_lambda_fm |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
844 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
845 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
846 |
lemma satisfies_is_d_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
847 |
"[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
848 |
u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
849 |
==> satisfies_is_d(**A,nu,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
850 |
sats(A, satisfies_is_d_fm(u,x,y,z), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
851 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
852 |
|
13494 | 853 |
theorem satisfies_is_d_reflection: |
854 |
"REFLECTS[\<lambda>x. satisfies_is_d(L,f(x),g(x),h(x),g'(x)), |
|
855 |
\<lambda>i x. satisfies_is_d(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
856 |
apply (unfold satisfies_is_d_def ) |
|
857 |
apply (intro FOL_reflections function_reflections is_lambda_reflection |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
858 |
extra_reflections nth_reflection depth_apply_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
859 |
is_list_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
860 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
861 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
862 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
863 |
subsubsection{*The Operator @{term satisfies_MH}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
864 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
865 |
(* satisfies_MH == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
866 |
\<lambda>M A u f zz. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
867 |
\<forall>fml[M]. is_formula(M,fml) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
868 |
is_lambda (M, fml, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
869 |
is_formula_case (M, satisfies_is_a(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
870 |
satisfies_is_b(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
871 |
satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
872 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
873 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
874 |
constdefs satisfies_MH_fm :: "[i,i,i,i]=>i" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
875 |
"satisfies_MH_fm(A,u,f,zz) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
876 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
877 |
Implies(is_formula_fm(0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
878 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
879 |
formula_case_fm(satisfies_is_a_fm(A#+7,2,1,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
880 |
satisfies_is_b_fm(A#+7,2,1,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
881 |
satisfies_is_c_fm(A#+7,f#+7,2,1,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
882 |
satisfies_is_d_fm(A#+6,f#+6,1,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
883 |
1, 0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
884 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
885 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
886 |
lemma satisfies_MH_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
887 |
"[| A \<in> nat; u \<in> nat; x \<in> nat; z \<in> nat |] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
888 |
==> satisfies_MH_fm(A,u,x,z) \<in> formula" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
889 |
by (simp add: satisfies_MH_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
890 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
891 |
lemma sats_satisfies_MH_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
892 |
"[| u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
893 |
==> sats(A, satisfies_MH_fm(u,x,y,z), env) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
894 |
satisfies_MH(**A, nth(u,env), nth(x,env), nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
895 |
by (simp add: satisfies_MH_fm_def satisfies_MH_def sats_lambda_fm |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
896 |
sats_formula_case_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
897 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
898 |
lemma satisfies_MH_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
899 |
"[| nth(u,env) = nu; nth(x,env) = nx; nth(y,env) = ny; nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
900 |
u \<in> nat; x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
901 |
==> satisfies_MH(**A,nu,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
902 |
sats(A, satisfies_MH_fm(u,x,y,z), env)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
903 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
904 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
905 |
lemmas satisfies_reflections = |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
906 |
is_lambda_reflection is_formula_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
907 |
is_formula_case_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
908 |
satisfies_is_a_reflection satisfies_is_b_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
909 |
satisfies_is_c_reflection satisfies_is_d_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
910 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
911 |
theorem satisfies_MH_reflection: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
912 |
"REFLECTS[\<lambda>x. satisfies_MH(L,f(x),g(x),h(x),g'(x)), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
913 |
\<lambda>i x. satisfies_MH(**Lset(i),f(x),g(x),h(x),g'(x))]" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
914 |
apply (unfold satisfies_MH_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
915 |
apply (intro FOL_reflections satisfies_reflections) |
13494 | 916 |
done |
917 |
||
918 |
||
13502 | 919 |
subsection{*Lemmas for Instantiating the Locale @{text "M_satisfies"}*} |
920 |
||
921 |
||
922 |
subsubsection{*The @{term "Member"} Case*} |
|
923 |
||
924 |
lemma Member_Reflects: |
|
925 |
"REFLECTS[\<lambda>u. \<exists>v[L]. v \<in> B \<and> (\<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
926 |
v \<in> lstA \<and> is_nth(L,x,v,nx) \<and> is_nth(L,y,v,ny) \<and> |
|
927 |
is_bool_of_o(L, nx \<in> ny, bo) \<and> pair(L,v,bo,u)), |
|
928 |
\<lambda>i u. \<exists>v \<in> Lset(i). v \<in> B \<and> (\<exists>bo \<in> Lset(i). \<exists>nx \<in> Lset(i). \<exists>ny \<in> Lset(i). |
|
929 |
v \<in> lstA \<and> is_nth(**Lset(i), x, v, nx) \<and> |
|
930 |
is_nth(**Lset(i), y, v, ny) \<and> |
|
931 |
is_bool_of_o(**Lset(i), nx \<in> ny, bo) \<and> pair(**Lset(i), v, bo, u))]" |
|
932 |
by (intro FOL_reflections function_reflections nth_reflection |
|
933 |
bool_of_o_reflection) |
|
934 |
||
935 |
||
936 |
lemma Member_replacement: |
|
937 |
"[|L(A); x \<in> nat; y \<in> nat|] |
|
938 |
==> strong_replacement |
|
939 |
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
940 |
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & |
|
941 |
is_bool_of_o(L, nx \<in> ny, bo) & |
|
942 |
pair(L, env, bo, z))" |
|
943 |
apply (frule list_closed) |
|
944 |
apply (rule strong_replacementI) |
|
945 |
apply (rule rallI) |
|
946 |
apply (rename_tac B) |
|
947 |
apply (rule separation_CollectI) |
|
948 |
apply (rule_tac A="{list(A),B,x,y,z}" in subset_LsetE, blast ) |
|
949 |
apply (rule ReflectsE [OF Member_Reflects], assumption) |
|
950 |
apply (drule subset_Lset_ltD, assumption) |
|
951 |
apply (erule reflection_imp_L_separation) |
|
952 |
apply (simp_all add: lt_Ord2) |
|
953 |
apply (simp add: is_nth_def is_wfrec_def is_bool_of_o_def) |
|
954 |
apply (rule DPow_LsetI) |
|
955 |
apply (rename_tac u) |
|
956 |
apply (rule bex_iff_sats conj_iff_sats conj_iff_sats)+ |
|
957 |
apply (rule_tac env = "[v,u,list(A),B,x,y,z]" in mem_iff_sats) |
|
958 |
apply (rule sep_rules is_nat_case_iff_sats iterates_MH_iff_sats |
|
959 |
is_recfun_iff_sats hd_iff_sats tl_iff_sats quasinat_iff_sats |
|
960 |
| simp)+ |
|
961 |
done |
|
962 |
||
963 |
||
964 |
subsubsection{*The @{term "Equal"} Case*} |
|
965 |
||
966 |
lemma Equal_Reflects: |
|
967 |
"REFLECTS[\<lambda>u. \<exists>v[L]. v \<in> B \<and> (\<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
968 |
v \<in> lstA \<and> is_nth(L, x, v, nx) \<and> is_nth(L, y, v, ny) \<and> |
|
969 |
is_bool_of_o(L, nx = ny, bo) \<and> pair(L, v, bo, u)), |
|
970 |
\<lambda>i u. \<exists>v \<in> Lset(i). v \<in> B \<and> (\<exists>bo \<in> Lset(i). \<exists>nx \<in> Lset(i). \<exists>ny \<in> Lset(i). |
|
971 |
v \<in> lstA \<and> is_nth(**Lset(i), x, v, nx) \<and> |
|
972 |
is_nth(**Lset(i), y, v, ny) \<and> |
|
973 |
is_bool_of_o(**Lset(i), nx = ny, bo) \<and> pair(**Lset(i), v, bo, u))]" |
|
974 |
by (intro FOL_reflections function_reflections nth_reflection |
|
975 |
bool_of_o_reflection) |
|
976 |
||
977 |
||
978 |
lemma Equal_replacement: |
|
979 |
"[|L(A); x \<in> nat; y \<in> nat|] |
|
980 |
==> strong_replacement |
|
981 |
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
982 |
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & |
|
983 |
is_bool_of_o(L, nx = ny, bo) & |
|
984 |
pair(L, env, bo, z))" |
|
985 |
apply (frule list_closed) |
|
986 |
apply (rule strong_replacementI) |
|
987 |
apply (rule rallI) |
|
988 |
apply (rename_tac B) |
|
989 |
apply (rule separation_CollectI) |
|
990 |
apply (rule_tac A="{list(A),B,x,y,z}" in subset_LsetE, blast ) |
|
991 |
apply (rule ReflectsE [OF Equal_Reflects], assumption) |
|
992 |
apply (drule subset_Lset_ltD, assumption) |
|
993 |
apply (erule reflection_imp_L_separation) |
|
994 |
apply (simp_all add: lt_Ord2) |
|
995 |
apply (simp add: is_nth_def is_wfrec_def is_bool_of_o_def) |
|
996 |
apply (rule DPow_LsetI) |
|
997 |
apply (rename_tac u) |
|
998 |
apply (rule bex_iff_sats conj_iff_sats conj_iff_sats)+ |
|
999 |
apply (rule_tac env = "[v,u,list(A),B,x,y,z]" in mem_iff_sats) |
|
1000 |
apply (rule sep_rules is_nat_case_iff_sats iterates_MH_iff_sats |
|
1001 |
is_recfun_iff_sats hd_iff_sats tl_iff_sats quasinat_iff_sats |
|
1002 |
| simp)+ |
|
1003 |
done |
|
1004 |
||
1005 |
subsubsection{*The @{term "Nand"} Case*} |
|
1006 |
||
1007 |
lemma Nand_Reflects: |
|
1008 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> |
|
1009 |
(\<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. |
|
1010 |
fun_apply(L, rp, u, rpe) \<and> fun_apply(L, rq, u, rqe) \<and> |
|
1011 |
is_and(L, rpe, rqe, andpq) \<and> is_not(L, andpq, notpq) \<and> |
|
1012 |
u \<in> list(A) \<and> pair(L, u, notpq, x)), |
|
1013 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> |
|
1014 |
(\<exists>rpe \<in> Lset(i). \<exists>rqe \<in> Lset(i). \<exists>andpq \<in> Lset(i). \<exists>notpq \<in> Lset(i). |
|
1015 |
fun_apply(**Lset(i), rp, u, rpe) \<and> fun_apply(**Lset(i), rq, u, rqe) \<and> |
|
1016 |
is_and(**Lset(i), rpe, rqe, andpq) \<and> is_not(**Lset(i), andpq, notpq) \<and> |
|
1017 |
u \<in> list(A) \<and> pair(**Lset(i), u, notpq, x))]" |
|
1018 |
apply (unfold is_and_def is_not_def) |
|
1019 |
apply (intro FOL_reflections function_reflections) |
|
1020 |
done |
|
1021 |
||
1022 |
lemma Nand_replacement: |
|
1023 |
"[|L(A); L(rp); L(rq)|] |
|
1024 |
==> strong_replacement |
|
1025 |
(L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. |
|
1026 |
fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) & |
|
1027 |
is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) & |
|
1028 |
env \<in> list(A) & pair(L, env, notpq, z))" |
|
1029 |
apply (frule list_closed) |
|
1030 |
apply (rule strong_replacementI) |
|
1031 |
apply (rule rallI) |
|
1032 |
apply (rename_tac B) |
|
1033 |
apply (rule separation_CollectI) |
|
1034 |
apply (rule_tac A="{list(A),B,rp,rq,z}" in subset_LsetE, blast ) |
|
1035 |
apply (rule ReflectsE [OF Nand_Reflects], assumption) |
|
1036 |
apply (drule subset_Lset_ltD, assumption) |
|
1037 |
apply (erule reflection_imp_L_separation) |
|
1038 |
apply (simp_all add: lt_Ord2) |
|
1039 |
apply (simp add: is_and_def is_not_def) |
|
1040 |
apply (rule DPow_LsetI) |
|
1041 |
apply (rename_tac v) |
|
1042 |
apply (rule bex_iff_sats conj_iff_sats conj_iff_sats)+ |
|
1043 |
apply (rule_tac env = "[u,v,list(A),B,rp,rq,z]" in mem_iff_sats) |
|
1044 |
apply (rule sep_rules | simp)+ |
|
1045 |
done |
|
1046 |
||
1047 |
||
1048 |
subsubsection{*The @{term "Forall"} Case*} |
|
1049 |
||
1050 |
lemma Forall_Reflects: |
|
1051 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>bo[L]. u \<in> list(A) \<and> |
|
1052 |
is_bool_of_o (L, |
|
1053 |
\<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. a \<in> A \<longrightarrow> |
|
1054 |
is_Cons(L,a,u,co) \<longrightarrow> fun_apply(L,rp,co,rpco) \<longrightarrow> |
|
1055 |
number1(L,rpco), |
|
1056 |
bo) \<and> pair(L,u,bo,x)), |
|
1057 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>bo \<in> Lset(i). u \<in> list(A) \<and> |
|
1058 |
is_bool_of_o (**Lset(i), |
|
1059 |
\<forall>a \<in> Lset(i). \<forall>co \<in> Lset(i). \<forall>rpco \<in> Lset(i). a \<in> A \<longrightarrow> |
|
1060 |
is_Cons(**Lset(i),a,u,co) \<longrightarrow> fun_apply(**Lset(i),rp,co,rpco) \<longrightarrow> |
|
1061 |
number1(**Lset(i),rpco), |
|
1062 |
bo) \<and> pair(**Lset(i),u,bo,x))]" |
|
1063 |
apply (unfold is_bool_of_o_def) |
|
1064 |
apply (intro FOL_reflections function_reflections Cons_reflection) |
|
1065 |
done |
|
1066 |
||
1067 |
lemma Forall_replacement: |
|
1068 |
"[|L(A); L(rp)|] |
|
1069 |
==> strong_replacement |
|
1070 |
(L, \<lambda>env z. \<exists>bo[L]. |
|
1071 |
env \<in> list(A) & |
|
1072 |
is_bool_of_o (L, |
|
1073 |
\<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. |
|
1074 |
a\<in>A --> is_Cons(L,a,env,co) --> |
|
1075 |
fun_apply(L,rp,co,rpco) --> number1(L, rpco), |
|
1076 |
bo) & |
|
1077 |
pair(L,env,bo,z))" |
|
1078 |
apply (frule list_closed) |
|
1079 |
apply (rule strong_replacementI) |
|
1080 |
apply (rule rallI) |
|
1081 |
apply (rename_tac B) |
|
1082 |
apply (rule separation_CollectI) |
|
1083 |
apply (rule_tac A="{A,list(A),B,rp,z}" in subset_LsetE, blast ) |
|
1084 |
apply (rule ReflectsE [OF Forall_Reflects], assumption) |
|
1085 |
apply (drule subset_Lset_ltD, assumption) |
|
1086 |
apply (erule reflection_imp_L_separation) |
|
1087 |
apply (simp_all add: lt_Ord2) |
|
1088 |
apply (simp add: is_bool_of_o_def) |
|
1089 |
apply (rule DPow_LsetI) |
|
1090 |
apply (rename_tac v) |
|
1091 |
apply (rule bex_iff_sats conj_iff_sats conj_iff_sats)+ |
|
1092 |
apply (rule_tac env = "[u,v,A,list(A),B,rp,z]" in mem_iff_sats) |
|
1093 |
apply (rule sep_rules Cons_iff_sats | simp)+ |
|
1094 |
done |
|
1095 |
||
1096 |
subsubsection{*The @{term "transrec_replacement"} Case*} |
|
1097 |
||
13494 | 1098 |
lemma formula_rec_replacement_Reflects: |
1099 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L, u, y, x) \<and> |
|
1100 |
is_wfrec (L, satisfies_MH(L,A), mesa, u, y)), |
|
1101 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
|
1102 |
is_wfrec (**Lset(i), satisfies_MH(**Lset(i),A), mesa, u, y))]" |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1103 |
by (intro FOL_reflections function_reflections satisfies_MH_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1104 |
is_wfrec_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1105 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1106 |
lemma formula_rec_replacement: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1107 |
--{*For the @{term transrec}*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1108 |
"[|n \<in> nat; L(A)|] ==> transrec_replacement(L, satisfies_MH(L,A), n)" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1109 |
apply (subgoal_tac "L(Memrel(eclose({n})))") |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1110 |
prefer 2 apply (simp add: nat_into_M) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1111 |
apply (rule transrec_replacementI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1112 |
apply (simp add: nat_into_M) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1113 |
apply (rule strong_replacementI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1114 |
apply (rule rallI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1115 |
apply (rename_tac B) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1116 |
apply (rule separation_CollectI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1117 |
apply (rule_tac A="{B,A,n,z,Memrel(eclose({n}))}" in subset_LsetE, blast ) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1118 |
apply (rule ReflectsE [OF formula_rec_replacement_Reflects], assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1119 |
apply (drule subset_Lset_ltD, assumption) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1120 |
apply (erule reflection_imp_L_separation) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1121 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1122 |
apply (elim conjE) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1123 |
apply (rule DPow_LsetI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1124 |
apply (rename_tac v) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1125 |
apply (rule bex_iff_sats conj_iff_sats)+ |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1126 |
apply (rule_tac env = "[u,v,A,n,B,Memrel(eclose({n}))]" in mem_iff_sats) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
1127 |
apply (rule sep_rules satisfies_MH_iff_sats is_wfrec_iff_sats | simp)+ |
13494 | 1128 |
done |
1129 |
||
13502 | 1130 |
|
1131 |
subsubsection{*The Lambda Replacement Case*} |
|
1132 |
||
1133 |
lemma formula_rec_lambda_replacement_Reflects: |
|
1134 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B & |
|
1135 |
mem_formula(L,u) & |
|
1136 |
(\<exists>c[L]. |
|
1137 |
is_formula_case |
|
1138 |
(L, satisfies_is_a(L,A), satisfies_is_b(L,A), |
|
1139 |
satisfies_is_c(L,A,g), satisfies_is_d(L,A,g), |
|
1140 |
u, c) & |
|
1141 |
pair(L,u,c,x)), |
|
1142 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & mem_formula(**Lset(i),u) & |
|
1143 |
(\<exists>c \<in> Lset(i). |
|
1144 |
is_formula_case |
|
1145 |
(**Lset(i), satisfies_is_a(**Lset(i),A), satisfies_is_b(**Lset(i),A), |
|
1146 |
satisfies_is_c(**Lset(i),A,g), satisfies_is_d(**Lset(i),A,g), |
|
1147 |
u, c) & |
|
1148 |
pair(**Lset(i),u,c,x))]" |
|
1149 |
by (intro FOL_reflections function_reflections mem_formula_reflection |
|
1150 |
is_formula_case_reflection satisfies_is_a_reflection |
|
1151 |
satisfies_is_b_reflection satisfies_is_c_reflection |
|
1152 |
satisfies_is_d_reflection) |
|
1153 |
||
1154 |
lemma formula_rec_lambda_replacement: |
|
1155 |
--{*For the @{term transrec}*} |
|
1156 |
"[|L(g); L(A)|] ==> |
|
1157 |
strong_replacement (L, |
|
1158 |
\<lambda>x y. mem_formula(L,x) & |
|
1159 |
(\<exists>c[L]. is_formula_case(L, satisfies_is_a(L,A), |
|
1160 |
satisfies_is_b(L,A), |
|
1161 |
satisfies_is_c(L,A,g), |
|
1162 |
satisfies_is_d(L,A,g), x, c) & |
|
1163 |
pair(L, x, c, y)))" |
|
1164 |
apply (rule strong_replacementI) |
|
1165 |
apply (rule rallI) |
|
1166 |
apply (rename_tac B) |
|
1167 |
apply (rule separation_CollectI) |
|
1168 |
apply (rule_tac A="{B,A,g,z}" in subset_LsetE, blast) |
|
1169 |
apply (rule ReflectsE [OF formula_rec_lambda_replacement_Reflects], assumption) |
|
1170 |
apply (drule subset_Lset_ltD, assumption) |
|
1171 |
apply (erule reflection_imp_L_separation) |
|
1172 |
apply (simp_all add: lt_Ord2 Memrel_closed) |
|
1173 |
apply (elim conjE) |
|
1174 |
apply (rule DPow_LsetI) |
|
1175 |
apply (rename_tac v) |
|
1176 |
apply (rule bex_iff_sats conj_iff_sats)+ |
|
1177 |
apply (rule_tac env = "[u,v,A,g,B]" in mem_iff_sats) |
|
1178 |
apply (rule sep_rules mem_formula_iff_sats |
|
1179 |
formula_case_iff_sats satisfies_is_a_iff_sats |
|
1180 |
satisfies_is_b_iff_sats satisfies_is_c_iff_sats |
|
1181 |
satisfies_is_d_iff_sats | simp)+ |
|
1182 |
done |
|
1183 |
||
1184 |
||
1185 |
subsection{*Instantiating @{text M_satisfies}*} |
|
1186 |
||
1187 |
lemma M_satisfies_axioms_L: "M_satisfies_axioms(L)" |
|
1188 |
apply (rule M_satisfies_axioms.intro) |
|
1189 |
apply (assumption | rule |
|
1190 |
Member_replacement Equal_replacement |
|
1191 |
Nand_replacement Forall_replacement |
|
1192 |
formula_rec_replacement formula_rec_lambda_replacement)+ |
|
1193 |
done |
|
1194 |
||
1195 |
theorem M_satisfies_L: "PROP M_satisfies(L)" |
|
1196 |
apply (rule M_satisfies.intro) |
|
1197 |
apply (rule M_eclose.axioms [OF M_eclose_L])+ |
|
1198 |
apply (rule M_satisfies_axioms_L) |
|
1199 |
done |
|
1200 |
||
13494 | 1201 |
end |