author | paulson |
Wed, 11 Sep 2002 16:55:37 +0200 | |
changeset 13566 | 52a419210d5c |
parent 13557 | 6061d0045409 |
child 13634 | 99a593b49b04 |
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 the Function @{term satisfies}*} |
|
175 |
||
176 |
constdefs |
|
177 |
is_depth_apply :: "[i=>o,i,i,i] => o" |
|
178 |
--{*Merely a useful abbreviation for the sequel.*} |
|
179 |
"is_depth_apply(M,h,p,z) == |
|
180 |
\<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
|
181 |
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) & |
13494 | 182 |
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)" |
183 |
||
184 |
lemma (in M_datatypes) is_depth_apply_abs [simp]: |
|
185 |
"[|M(h); p \<in> formula; M(z)|] |
|
186 |
==> is_depth_apply(M,h,p,z) <-> z = h ` succ(depth(p)) ` p" |
|
187 |
by (simp add: is_depth_apply_def formula_into_M depth_type eq_commute) |
|
188 |
||
189 |
||
190 |
||
191 |
text{*There is at present some redundancy between the relativizations in |
|
192 |
e.g. @{text satisfies_is_a} and those in e.g. @{text Member_replacement}.*} |
|
193 |
||
194 |
text{*These constants let us instantiate the parameters @{term a}, @{term b}, |
|
13504 | 195 |
@{term c}, @{term d}, etc., of the locale @{text Formula_Rec}.*} |
13494 | 196 |
constdefs |
197 |
satisfies_a :: "[i,i,i]=>i" |
|
198 |
"satisfies_a(A) == |
|
199 |
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env))" |
|
200 |
||
201 |
satisfies_is_a :: "[i=>o,i,i,i,i]=>o" |
|
202 |
"satisfies_is_a(M,A) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
203 |
\<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
|
204 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
205 |
\<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
|
206 |
\<exists>nx[M]. \<exists>ny[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
207 |
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
|
208 |
zz)" |
13494 | 209 |
|
210 |
satisfies_b :: "[i,i,i]=>i" |
|
211 |
"satisfies_b(A) == |
|
212 |
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env))" |
|
213 |
||
214 |
satisfies_is_b :: "[i=>o,i,i,i,i]=>o" |
|
215 |
--{*We simplify the formula to have just @{term nx} rather than |
|
216 |
introducing @{term ny} with @{term "nx=ny"} *} |
|
217 |
"satisfies_is_b(M,A) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
218 |
\<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
|
219 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
220 |
\<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
|
221 |
\<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
|
222 |
zz)" |
13494 | 223 |
|
224 |
satisfies_c :: "[i,i,i,i,i]=>i" |
|
13502 | 225 |
"satisfies_c(A) == \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)" |
13494 | 226 |
|
227 |
satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" |
|
228 |
"satisfies_is_c(M,A,h) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
229 |
\<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
|
230 |
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. |
13494 | 231 |
(\<exists>rp[M]. is_depth_apply(M,h,p,rp) & fun_apply(M,rp,env,hp)) & |
232 |
(\<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
|
233 |
(\<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
|
234 |
zz)" |
13494 | 235 |
|
236 |
satisfies_d :: "[i,i,i]=>i" |
|
237 |
"satisfies_d(A) |
|
238 |
== \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)" |
|
239 |
||
240 |
satisfies_is_d :: "[i=>o,i,i,i,i]=>o" |
|
241 |
"satisfies_is_d(M,A,h) == |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
242 |
\<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
|
243 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
244 |
\<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
|
245 |
is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
246 |
\<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
|
247 |
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
|
248 |
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
|
249 |
z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
250 |
zz)" |
13494 | 251 |
|
252 |
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
|
253 |
--{*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
|
254 |
the correct arity.*} |
13494 | 255 |
"satisfies_MH == |
13502 | 256 |
\<lambda>M A u f z. |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
257 |
\<forall>fml[M]. is_formula(M,fml) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
258 |
is_lambda (M, fml, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
259 |
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
|
260 |
satisfies_is_b(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
261 |
satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)), |
13502 | 262 |
z)" |
13494 | 263 |
|
13502 | 264 |
is_satisfies :: "[i=>o,i,i,i]=>o" |
13503 | 265 |
"is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))" |
13502 | 266 |
|
267 |
||
268 |
text{*This lemma relates the fragments defined above to the original primitive |
|
269 |
recursion in @{term satisfies}. |
|
270 |
Induction is not required: the definitions are directly equal!*} |
|
271 |
lemma satisfies_eq: |
|
272 |
"satisfies(A,p) = |
|
273 |
formula_rec (satisfies_a(A), satisfies_b(A), |
|
274 |
satisfies_c(A), satisfies_d(A), p)" |
|
275 |
by (simp add: satisfies_formula_def satisfies_a_def satisfies_b_def |
|
276 |
satisfies_c_def satisfies_d_def) |
|
13494 | 277 |
|
278 |
text{*Further constraints on the class @{term M} in order to prove |
|
279 |
absoluteness for the constants defined above. The ultimate goal |
|
280 |
is the absoluteness of the function @{term satisfies}. *} |
|
13502 | 281 |
locale M_satisfies = M_eclose + |
13494 | 282 |
assumes |
283 |
Member_replacement: |
|
284 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
285 |
==> strong_replacement |
|
286 |
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. |
|
287 |
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & |
|
288 |
is_bool_of_o(M, nx \<in> ny, bo) & |
|
289 |
pair(M, env, bo, z))" |
|
290 |
and |
|
291 |
Equal_replacement: |
|
292 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
293 |
==> strong_replacement |
|
294 |
(M, \<lambda>env z. \<exists>bo[M]. \<exists>nx[M]. \<exists>ny[M]. |
|
295 |
env \<in> list(A) & is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & |
|
296 |
is_bool_of_o(M, nx = ny, bo) & |
|
297 |
pair(M, env, bo, z))" |
|
298 |
and |
|
299 |
Nand_replacement: |
|
300 |
"[|M(A); M(rp); M(rq)|] |
|
301 |
==> strong_replacement |
|
302 |
(M, \<lambda>env z. \<exists>rpe[M]. \<exists>rqe[M]. \<exists>andpq[M]. \<exists>notpq[M]. |
|
303 |
fun_apply(M,rp,env,rpe) & fun_apply(M,rq,env,rqe) & |
|
304 |
is_and(M,rpe,rqe,andpq) & is_not(M,andpq,notpq) & |
|
305 |
env \<in> list(A) & pair(M, env, notpq, z))" |
|
306 |
and |
|
307 |
Forall_replacement: |
|
308 |
"[|M(A); M(rp)|] |
|
309 |
==> strong_replacement |
|
310 |
(M, \<lambda>env z. \<exists>bo[M]. |
|
311 |
env \<in> list(A) & |
|
312 |
is_bool_of_o (M, |
|
313 |
\<forall>a[M]. \<forall>co[M]. \<forall>rpco[M]. |
|
314 |
a\<in>A --> is_Cons(M,a,env,co) --> |
|
315 |
fun_apply(M,rp,co,rpco) --> number1(M, rpco), |
|
316 |
bo) & |
|
317 |
pair(M,env,bo,z))" |
|
318 |
and |
|
319 |
formula_rec_replacement: |
|
320 |
--{*For the @{term transrec}*} |
|
321 |
"[|n \<in> nat; M(A)|] ==> transrec_replacement(M, satisfies_MH(M,A), n)" |
|
322 |
and |
|
323 |
formula_rec_lambda_replacement: |
|
324 |
--{*For the @{text "\<lambda>-abstraction"} in the @{term transrec} body*} |
|
13502 | 325 |
"[|M(g); M(A)|] ==> |
326 |
strong_replacement (M, |
|
327 |
\<lambda>x y. mem_formula(M,x) & |
|
328 |
(\<exists>c[M]. is_formula_case(M, satisfies_is_a(M,A), |
|
329 |
satisfies_is_b(M,A), |
|
330 |
satisfies_is_c(M,A,g), |
|
331 |
satisfies_is_d(M,A,g), x, c) & |
|
332 |
pair(M, x, c, y)))" |
|
13494 | 333 |
|
334 |
||
335 |
lemma (in M_satisfies) Member_replacement': |
|
336 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
337 |
==> strong_replacement |
|
338 |
(M, \<lambda>env z. env \<in> list(A) & |
|
339 |
z = \<langle>env, bool_of_o(nth(x, env) \<in> nth(y, env))\<rangle>)" |
|
340 |
by (insert Member_replacement, simp) |
|
341 |
||
342 |
lemma (in M_satisfies) Equal_replacement': |
|
343 |
"[|M(A); x \<in> nat; y \<in> nat|] |
|
344 |
==> strong_replacement |
|
345 |
(M, \<lambda>env z. env \<in> list(A) & |
|
346 |
z = \<langle>env, bool_of_o(nth(x, env) = nth(y, env))\<rangle>)" |
|
347 |
by (insert Equal_replacement, simp) |
|
348 |
||
349 |
lemma (in M_satisfies) Nand_replacement': |
|
350 |
"[|M(A); M(rp); M(rq)|] |
|
351 |
==> strong_replacement |
|
352 |
(M, \<lambda>env z. env \<in> list(A) & z = \<langle>env, not(rp`env and rq`env)\<rangle>)" |
|
353 |
by (insert Nand_replacement, simp) |
|
354 |
||
355 |
lemma (in M_satisfies) Forall_replacement': |
|
356 |
"[|M(A); M(rp)|] |
|
357 |
==> strong_replacement |
|
358 |
(M, \<lambda>env z. |
|
359 |
env \<in> list(A) & |
|
360 |
z = \<langle>env, bool_of_o (\<forall>a\<in>A. rp ` Cons(a,env) = 1)\<rangle>)" |
|
361 |
by (insert Forall_replacement, simp) |
|
362 |
||
363 |
lemma (in M_satisfies) a_closed: |
|
364 |
"[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_a(A,x,y))" |
|
365 |
apply (simp add: satisfies_a_def) |
|
366 |
apply (blast intro: lam_closed2 Member_replacement') |
|
367 |
done |
|
368 |
||
369 |
lemma (in M_satisfies) a_rel: |
|
370 |
"M(A) ==> Relativize2(M, nat, nat, satisfies_is_a(M,A), satisfies_a(A))" |
|
371 |
apply (simp add: Relativize2_def satisfies_is_a_def satisfies_a_def) |
|
372 |
apply (simp add: lambda_abs2 [OF Member_replacement'] Relativize1_def) |
|
373 |
done |
|
374 |
||
375 |
lemma (in M_satisfies) b_closed: |
|
376 |
"[|M(A); x\<in>nat; y\<in>nat|] ==> M(satisfies_b(A,x,y))" |
|
377 |
apply (simp add: satisfies_b_def) |
|
378 |
apply (blast intro: lam_closed2 Equal_replacement') |
|
379 |
done |
|
380 |
||
381 |
lemma (in M_satisfies) b_rel: |
|
382 |
"M(A) ==> Relativize2(M, nat, nat, satisfies_is_b(M,A), satisfies_b(A))" |
|
383 |
apply (simp add: Relativize2_def satisfies_is_b_def satisfies_b_def) |
|
384 |
apply (simp add: lambda_abs2 [OF Equal_replacement'] Relativize1_def) |
|
385 |
done |
|
386 |
||
387 |
lemma (in M_satisfies) c_closed: |
|
388 |
"[|M(A); x \<in> formula; y \<in> formula; M(rx); M(ry)|] |
|
389 |
==> M(satisfies_c(A,x,y,rx,ry))" |
|
390 |
apply (simp add: satisfies_c_def) |
|
391 |
apply (rule lam_closed2) |
|
392 |
apply (rule Nand_replacement') |
|
393 |
apply (simp_all add: formula_into_M list_into_M [of _ A]) |
|
394 |
done |
|
395 |
||
396 |
lemma (in M_satisfies) c_rel: |
|
397 |
"[|M(A); M(f)|] ==> |
|
398 |
Relativize2 (M, formula, formula, |
|
399 |
satisfies_is_c(M,A,f), |
|
400 |
\<lambda>u v. satisfies_c(A, u, v, f ` succ(depth(u)) ` u, |
|
401 |
f ` succ(depth(v)) ` v))" |
|
402 |
apply (simp add: Relativize2_def satisfies_is_c_def satisfies_c_def) |
|
403 |
apply (simp add: lambda_abs2 [OF Nand_replacement' triv_Relativize1] |
|
404 |
formula_into_M) |
|
405 |
done |
|
406 |
||
407 |
lemma (in M_satisfies) d_closed: |
|
408 |
"[|M(A); x \<in> formula; M(rx)|] ==> M(satisfies_d(A,x,rx))" |
|
409 |
apply (simp add: satisfies_d_def) |
|
410 |
apply (rule lam_closed2) |
|
411 |
apply (rule Forall_replacement') |
|
412 |
apply (simp_all add: formula_into_M list_into_M [of _ A]) |
|
413 |
done |
|
414 |
||
415 |
lemma (in M_satisfies) d_rel: |
|
416 |
"[|M(A); M(f)|] ==> |
|
417 |
Relativize1(M, formula, satisfies_is_d(M,A,f), |
|
418 |
\<lambda>u. satisfies_d(A, u, f ` succ(depth(u)) ` u))" |
|
419 |
apply (simp del: rall_abs |
|
420 |
add: Relativize1_def satisfies_is_d_def satisfies_d_def) |
|
421 |
apply (simp add: lambda_abs2 [OF Forall_replacement' triv_Relativize1] |
|
422 |
formula_into_M) |
|
423 |
done |
|
424 |
||
425 |
||
426 |
lemma (in M_satisfies) fr_replace: |
|
427 |
"[|n \<in> nat; M(A)|] ==> transrec_replacement(M,satisfies_MH(M,A),n)" |
|
428 |
by (blast intro: formula_rec_replacement) |
|
429 |
||
13502 | 430 |
lemma (in M_satisfies) formula_case_satisfies_closed: |
431 |
"[|M(g); M(A); x \<in> formula|] ==> |
|
432 |
M(formula_case (satisfies_a(A), satisfies_b(A), |
|
433 |
\<lambda>u v. satisfies_c(A, u, v, |
|
434 |
g ` succ(depth(u)) ` u, g ` succ(depth(v)) ` v), |
|
435 |
\<lambda>u. satisfies_d (A, u, g ` succ(depth(u)) ` u), |
|
436 |
x))" |
|
437 |
by (blast intro: formula_case_closed a_closed b_closed c_closed d_closed) |
|
438 |
||
13494 | 439 |
lemma (in M_satisfies) fr_lam_replace: |
13502 | 440 |
"[|M(g); M(A)|] ==> |
13494 | 441 |
strong_replacement (M, \<lambda>x y. x \<in> formula & |
442 |
y = \<langle>x, |
|
443 |
formula_rec_case(satisfies_a(A), |
|
444 |
satisfies_b(A), |
|
445 |
satisfies_c(A), |
|
446 |
satisfies_d(A), g, x)\<rangle>)" |
|
13502 | 447 |
apply (insert formula_rec_lambda_replacement) |
448 |
apply (simp add: formula_rec_case_def formula_case_satisfies_closed |
|
449 |
formula_case_abs [OF a_rel b_rel c_rel d_rel]) |
|
450 |
done |
|
13494 | 451 |
|
452 |
||
453 |
||
13504 | 454 |
text{*Instantiate locale @{text Formula_Rec} for the |
13502 | 455 |
Function @{term satisfies}*} |
13494 | 456 |
|
13504 | 457 |
lemma (in M_satisfies) Formula_Rec_axioms_M: |
13502 | 458 |
"M(A) ==> |
13504 | 459 |
Formula_Rec_axioms(M, satisfies_a(A), satisfies_is_a(M,A), |
460 |
satisfies_b(A), satisfies_is_b(M,A), |
|
461 |
satisfies_c(A), satisfies_is_c(M,A), |
|
462 |
satisfies_d(A), satisfies_is_d(M,A))" |
|
463 |
apply (rule Formula_Rec_axioms.intro) |
|
13502 | 464 |
apply (assumption | |
465 |
rule a_closed a_rel b_closed b_rel c_closed c_rel d_closed d_rel |
|
466 |
fr_replace [unfolded satisfies_MH_def] |
|
467 |
fr_lam_replace) + |
|
13494 | 468 |
done |
469 |
||
470 |
||
13504 | 471 |
theorem (in M_satisfies) Formula_Rec_M: |
13502 | 472 |
"M(A) ==> |
13504 | 473 |
PROP Formula_Rec(M, satisfies_a(A), satisfies_is_a(M,A), |
474 |
satisfies_b(A), satisfies_is_b(M,A), |
|
475 |
satisfies_c(A), satisfies_is_c(M,A), |
|
476 |
satisfies_d(A), satisfies_is_d(M,A))" |
|
477 |
apply (rule Formula_Rec.intro, assumption+) |
|
478 |
apply (erule Formula_Rec_axioms_M) |
|
13494 | 479 |
done |
480 |
||
13502 | 481 |
lemmas (in M_satisfies) |
13535 | 482 |
satisfies_closed' = Formula_Rec.formula_rec_closed [OF Formula_Rec_M] |
483 |
and satisfies_abs' = Formula_Rec.formula_rec_abs [OF Formula_Rec_M] |
|
13494 | 484 |
|
485 |
||
13502 | 486 |
lemma (in M_satisfies) satisfies_closed: |
487 |
"[|M(A); p \<in> formula|] ==> M(satisfies(A,p))" |
|
13504 | 488 |
by (simp add: Formula_Rec.formula_rec_closed [OF Formula_Rec_M] |
13502 | 489 |
satisfies_eq) |
13494 | 490 |
|
13502 | 491 |
lemma (in M_satisfies) satisfies_abs: |
492 |
"[|M(A); M(z); p \<in> formula|] |
|
493 |
==> is_satisfies(M,A,p,z) <-> z = satisfies(A,p)" |
|
13504 | 494 |
by (simp only: Formula_Rec.formula_rec_abs [OF Formula_Rec_M] |
13503 | 495 |
satisfies_eq is_satisfies_def satisfies_MH_def) |
13494 | 496 |
|
497 |
||
13502 | 498 |
subsection{*Internalizations Needed to Instantiate @{text "M_satisfies"}*} |
13494 | 499 |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
500 |
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
|
501 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
502 |
(* is_depth_apply(M,h,p,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
503 |
\<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
|
504 |
2 1 0 |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
505 |
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
|
506 |
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
|
507 |
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
|
508 |
"depth_apply_fm(h,p,z) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
509 |
Exists(Exists(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
510 |
And(finite_ordinal_fm(2), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
511 |
And(depth_fm(p#+3,2), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
512 |
And(succ_fm(2,1), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
513 |
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
|
514 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
515 |
lemma depth_apply_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
516 |
"[| 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
|
517 |
by (simp add: depth_apply_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
518 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
519 |
lemma sats_depth_apply_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
520 |
"[| 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
|
521 |
==> 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
|
522 |
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
|
523 |
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
|
524 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
525 |
lemma depth_apply_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
526 |
"[| 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
|
527 |
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
|
528 |
==> 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
|
529 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
530 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
531 |
lemma depth_apply_reflection: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
532 |
"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
|
533 |
\<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
|
534 |
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
|
535 |
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
|
536 |
finite_ordinal_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
537 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
538 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
539 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
540 |
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
|
541 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
542 |
(* satisfies_is_a(M,A) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
543 |
\<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
|
544 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
545 |
\<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
|
546 |
\<exists>nx[M]. \<exists>ny[M]. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
547 |
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
|
548 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
549 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
550 |
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
|
551 |
"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
|
552 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
553 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
554 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
555 |
bool_of_o_fm(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
556 |
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
|
557 |
And(nth_fm(y#+6,3,0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
558 |
Member(1,0))))), 0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
559 |
0, succ(z))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
560 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
561 |
lemma satisfies_is_a_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
562 |
"[| 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
|
563 |
==> 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
|
564 |
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
|
565 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
566 |
lemma sats_satisfies_is_a_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
567 |
"[| 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
|
568 |
==> 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
|
569 |
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
|
570 |
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
|
571 |
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
|
572 |
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
|
573 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
574 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
575 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
576 |
lemma satisfies_is_a_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
577 |
"[| 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
|
578 |
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
|
579 |
==> 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
|
580 |
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
|
581 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
582 |
|
13494 | 583 |
theorem satisfies_is_a_reflection: |
584 |
"REFLECTS[\<lambda>x. satisfies_is_a(L,f(x),g(x),h(x),g'(x)), |
|
585 |
\<lambda>i x. satisfies_is_a(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
586 |
apply (unfold satisfies_is_a_def) |
|
587 |
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
|
588 |
nth_reflection is_list_reflection) |
13494 | 589 |
done |
590 |
||
591 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
592 |
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
|
593 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
594 |
(* satisfies_is_b(M,A) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
595 |
\<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
|
596 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
597 |
\<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
|
598 |
\<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
|
599 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
600 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
601 |
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
|
602 |
"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
|
603 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
604 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
605 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
606 |
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
|
607 |
0, succ(z))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
608 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
609 |
lemma satisfies_is_b_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
610 |
"[| 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
|
611 |
==> 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
|
612 |
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
|
613 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
614 |
lemma sats_satisfies_is_b_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
615 |
"[| 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
|
616 |
==> 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
|
617 |
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
|
618 |
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
|
619 |
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
|
620 |
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
|
621 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
622 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
623 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
624 |
lemma satisfies_is_b_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
625 |
"[| 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
|
626 |
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
|
627 |
==> 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
|
628 |
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
|
629 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
630 |
|
13494 | 631 |
theorem satisfies_is_b_reflection: |
632 |
"REFLECTS[\<lambda>x. satisfies_is_b(L,f(x),g(x),h(x),g'(x)), |
|
633 |
\<lambda>i x. satisfies_is_b(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
634 |
apply (unfold satisfies_is_b_def) |
|
635 |
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
|
636 |
nth_reflection is_list_reflection) |
13494 | 637 |
done |
638 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
639 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
640 |
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
|
641 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
642 |
(* satisfies_is_c(M,A,h) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
643 |
\<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
|
644 |
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
|
645 |
(\<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
|
646 |
(\<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
|
647 |
(\<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
|
648 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
649 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
650 |
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
|
651 |
"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
|
652 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
653 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
654 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
655 |
Exists(Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
656 |
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
|
657 |
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
|
658 |
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
|
659 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
660 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
661 |
lemma satisfies_is_c_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
662 |
"[| 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
|
663 |
==> 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
|
664 |
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
|
665 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
666 |
lemma sats_satisfies_is_c_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
667 |
"[| 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
|
668 |
==> 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
|
669 |
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
|
670 |
nth(y,env), nth(z,env))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
671 |
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
|
672 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
673 |
lemma satisfies_is_c_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
674 |
"[| 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
|
675 |
nth(z,env) = nz; |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
676 |
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
|
677 |
==> 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
|
678 |
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
|
679 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
680 |
|
13494 | 681 |
theorem satisfies_is_c_reflection: |
682 |
"REFLECTS[\<lambda>x. satisfies_is_c(L,f(x),g(x),h(x),g'(x),h'(x)), |
|
683 |
\<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
|
684 |
apply (unfold satisfies_is_c_def) |
13494 | 685 |
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
|
686 |
extra_reflections nth_reflection depth_apply_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
687 |
is_list_reflection) |
13494 | 688 |
done |
689 |
||
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
690 |
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
|
691 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
692 |
(* satisfies_is_d(M,A,h) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
693 |
\<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
|
694 |
is_lambda(M, lA, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
695 |
\<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
|
696 |
is_bool_of_o(M, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
697 |
\<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
|
698 |
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
|
699 |
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
|
700 |
z), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
701 |
zz) *) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
702 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
703 |
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
|
704 |
"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
|
705 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
706 |
Implies(is_list_fm(succ(A),0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
707 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
708 |
Exists( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
709 |
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
|
710 |
bool_of_o_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
711 |
Forall(Forall(Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
712 |
Implies(Member(2,A#+8), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
713 |
Implies(Cons_fm(2,5,1), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
714 |
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
|
715 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
716 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
717 |
lemma satisfies_is_d_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
718 |
"[| 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
|
719 |
==> 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
|
720 |
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
|
721 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
722 |
lemma sats_satisfies_is_d_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
723 |
"[| 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
|
724 |
==> 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
|
725 |
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
|
726 |
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
|
727 |
sats_bool_of_o_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
728 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
729 |
lemma satisfies_is_d_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
730 |
"[| 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
|
731 |
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
|
732 |
==> 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
|
733 |
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
|
734 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
735 |
|
13494 | 736 |
theorem satisfies_is_d_reflection: |
737 |
"REFLECTS[\<lambda>x. satisfies_is_d(L,f(x),g(x),h(x),g'(x)), |
|
738 |
\<lambda>i x. satisfies_is_d(**Lset(i),f(x),g(x),h(x),g'(x))]" |
|
13505 | 739 |
apply (unfold satisfies_is_d_def) |
13494 | 740 |
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
|
741 |
extra_reflections nth_reflection depth_apply_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
742 |
is_list_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
743 |
done |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
744 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
745 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
746 |
subsubsection{*The Operator @{term satisfies_MH}, Internalized*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
747 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
748 |
(* satisfies_MH == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
749 |
\<lambda>M A u f zz. |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
750 |
\<forall>fml[M]. is_formula(M,fml) --> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
751 |
is_lambda (M, fml, |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
752 |
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
|
753 |
satisfies_is_b(M,A), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
754 |
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
|
755 |
zz) *) |
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 |
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
|
758 |
"satisfies_MH_fm(A,u,f,zz) == |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
759 |
Forall( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
760 |
Implies(is_formula_fm(0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
761 |
lambda_fm( |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
762 |
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
|
763 |
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
|
764 |
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
|
765 |
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
|
766 |
1, 0), |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
767 |
0, succ(zz))))" |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
768 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
769 |
lemma satisfies_MH_type [TC]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
770 |
"[| 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
|
771 |
==> 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
|
772 |
by (simp add: satisfies_MH_fm_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
773 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
774 |
lemma sats_satisfies_MH_fm [simp]: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
775 |
"[| 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
|
776 |
==> 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
|
777 |
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
|
778 |
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
|
779 |
sats_formula_case_fm) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
780 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
781 |
lemma satisfies_MH_iff_sats: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
782 |
"[| 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
|
783 |
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
|
784 |
==> satisfies_MH(**A,nu,nx,ny,nz) <-> |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
785 |
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
|
786 |
by simp |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
787 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
788 |
lemmas satisfies_reflections = |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
789 |
is_lambda_reflection is_formula_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
790 |
is_formula_case_reflection |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
791 |
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
|
792 |
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
|
793 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
794 |
theorem satisfies_MH_reflection: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
795 |
"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
|
796 |
\<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
|
797 |
apply (unfold satisfies_MH_def) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
798 |
apply (intro FOL_reflections satisfies_reflections) |
13494 | 799 |
done |
800 |
||
801 |
||
13502 | 802 |
subsection{*Lemmas for Instantiating the Locale @{text "M_satisfies"}*} |
803 |
||
804 |
||
805 |
subsubsection{*The @{term "Member"} Case*} |
|
806 |
||
807 |
lemma Member_Reflects: |
|
808 |
"REFLECTS[\<lambda>u. \<exists>v[L]. v \<in> B \<and> (\<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
809 |
v \<in> lstA \<and> is_nth(L,x,v,nx) \<and> is_nth(L,y,v,ny) \<and> |
|
810 |
is_bool_of_o(L, nx \<in> ny, bo) \<and> pair(L,v,bo,u)), |
|
811 |
\<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). |
|
812 |
v \<in> lstA \<and> is_nth(**Lset(i), x, v, nx) \<and> |
|
813 |
is_nth(**Lset(i), y, v, ny) \<and> |
|
814 |
is_bool_of_o(**Lset(i), nx \<in> ny, bo) \<and> pair(**Lset(i), v, bo, u))]" |
|
815 |
by (intro FOL_reflections function_reflections nth_reflection |
|
816 |
bool_of_o_reflection) |
|
817 |
||
818 |
||
819 |
lemma Member_replacement: |
|
820 |
"[|L(A); x \<in> nat; y \<in> nat|] |
|
821 |
==> strong_replacement |
|
822 |
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
823 |
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & |
|
824 |
is_bool_of_o(L, nx \<in> ny, bo) & |
|
825 |
pair(L, env, bo, z))" |
|
13566 | 826 |
apply (rule strong_replacementI) |
827 |
apply (rename_tac B) |
|
828 |
apply (rule_tac u="{list(A),B,x,y}" |
|
829 |
in gen_separation [OF Member_Reflects], |
|
830 |
simp add: nat_into_M list_closed) |
|
831 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13502 | 832 |
apply (rule DPow_LsetI) |
13566 | 833 |
apply (rule bex_iff_sats conj_iff_sats)+ |
834 |
apply (rule_tac env = "[v,u,list(A),B,x,y]" in mem_iff_sats) |
|
835 |
apply (rule sep_rules nth_iff_sats is_bool_of_o_iff_sats | simp)+ |
|
13502 | 836 |
done |
837 |
||
838 |
||
839 |
subsubsection{*The @{term "Equal"} Case*} |
|
840 |
||
841 |
lemma Equal_Reflects: |
|
842 |
"REFLECTS[\<lambda>u. \<exists>v[L]. v \<in> B \<and> (\<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
843 |
v \<in> lstA \<and> is_nth(L, x, v, nx) \<and> is_nth(L, y, v, ny) \<and> |
|
844 |
is_bool_of_o(L, nx = ny, bo) \<and> pair(L, v, bo, u)), |
|
845 |
\<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). |
|
846 |
v \<in> lstA \<and> is_nth(**Lset(i), x, v, nx) \<and> |
|
847 |
is_nth(**Lset(i), y, v, ny) \<and> |
|
848 |
is_bool_of_o(**Lset(i), nx = ny, bo) \<and> pair(**Lset(i), v, bo, u))]" |
|
849 |
by (intro FOL_reflections function_reflections nth_reflection |
|
850 |
bool_of_o_reflection) |
|
851 |
||
852 |
||
853 |
lemma Equal_replacement: |
|
854 |
"[|L(A); x \<in> nat; y \<in> nat|] |
|
855 |
==> strong_replacement |
|
856 |
(L, \<lambda>env z. \<exists>bo[L]. \<exists>nx[L]. \<exists>ny[L]. |
|
857 |
env \<in> list(A) & is_nth(L,x,env,nx) & is_nth(L,y,env,ny) & |
|
858 |
is_bool_of_o(L, nx = ny, bo) & |
|
859 |
pair(L, env, bo, z))" |
|
13566 | 860 |
apply (rule strong_replacementI) |
861 |
apply (rename_tac B) |
|
862 |
apply (rule_tac u="{list(A),B,x,y}" |
|
863 |
in gen_separation [OF Equal_Reflects], |
|
864 |
simp add: nat_into_M list_closed) |
|
865 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13502 | 866 |
apply (rule DPow_LsetI) |
13566 | 867 |
apply (rule bex_iff_sats conj_iff_sats)+ |
868 |
apply (rule_tac env = "[v,u,list(A),B,x,y]" in mem_iff_sats) |
|
869 |
apply (rule sep_rules nth_iff_sats is_bool_of_o_iff_sats | simp)+ |
|
13502 | 870 |
done |
871 |
||
872 |
subsubsection{*The @{term "Nand"} Case*} |
|
873 |
||
874 |
lemma Nand_Reflects: |
|
875 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> |
|
876 |
(\<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. |
|
877 |
fun_apply(L, rp, u, rpe) \<and> fun_apply(L, rq, u, rqe) \<and> |
|
878 |
is_and(L, rpe, rqe, andpq) \<and> is_not(L, andpq, notpq) \<and> |
|
879 |
u \<in> list(A) \<and> pair(L, u, notpq, x)), |
|
880 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> |
|
881 |
(\<exists>rpe \<in> Lset(i). \<exists>rqe \<in> Lset(i). \<exists>andpq \<in> Lset(i). \<exists>notpq \<in> Lset(i). |
|
882 |
fun_apply(**Lset(i), rp, u, rpe) \<and> fun_apply(**Lset(i), rq, u, rqe) \<and> |
|
883 |
is_and(**Lset(i), rpe, rqe, andpq) \<and> is_not(**Lset(i), andpq, notpq) \<and> |
|
884 |
u \<in> list(A) \<and> pair(**Lset(i), u, notpq, x))]" |
|
885 |
apply (unfold is_and_def is_not_def) |
|
886 |
apply (intro FOL_reflections function_reflections) |
|
887 |
done |
|
888 |
||
889 |
lemma Nand_replacement: |
|
890 |
"[|L(A); L(rp); L(rq)|] |
|
891 |
==> strong_replacement |
|
892 |
(L, \<lambda>env z. \<exists>rpe[L]. \<exists>rqe[L]. \<exists>andpq[L]. \<exists>notpq[L]. |
|
893 |
fun_apply(L,rp,env,rpe) & fun_apply(L,rq,env,rqe) & |
|
894 |
is_and(L,rpe,rqe,andpq) & is_not(L,andpq,notpq) & |
|
895 |
env \<in> list(A) & pair(L, env, notpq, z))" |
|
13566 | 896 |
apply (rule strong_replacementI) |
897 |
apply (rename_tac B) |
|
898 |
apply (rule_tac u="{list(A),B,rp,rq}" in gen_separation [OF Nand_Reflects], |
|
899 |
simp add: list_closed) |
|
900 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13502 | 901 |
apply (rule DPow_LsetI) |
13566 | 902 |
apply (rule bex_iff_sats conj_iff_sats)+ |
903 |
apply (rule_tac env = "[u,x,list(A),B,rp,rq]" in mem_iff_sats) |
|
904 |
apply (rule sep_rules is_and_iff_sats is_not_iff_sats | simp)+ |
|
13502 | 905 |
done |
906 |
||
907 |
||
908 |
subsubsection{*The @{term "Forall"} Case*} |
|
909 |
||
910 |
lemma Forall_Reflects: |
|
911 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>bo[L]. u \<in> list(A) \<and> |
|
912 |
is_bool_of_o (L, |
|
913 |
\<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. a \<in> A \<longrightarrow> |
|
914 |
is_Cons(L,a,u,co) \<longrightarrow> fun_apply(L,rp,co,rpco) \<longrightarrow> |
|
915 |
number1(L,rpco), |
|
916 |
bo) \<and> pair(L,u,bo,x)), |
|
917 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>bo \<in> Lset(i). u \<in> list(A) \<and> |
|
918 |
is_bool_of_o (**Lset(i), |
|
919 |
\<forall>a \<in> Lset(i). \<forall>co \<in> Lset(i). \<forall>rpco \<in> Lset(i). a \<in> A \<longrightarrow> |
|
920 |
is_Cons(**Lset(i),a,u,co) \<longrightarrow> fun_apply(**Lset(i),rp,co,rpco) \<longrightarrow> |
|
921 |
number1(**Lset(i),rpco), |
|
922 |
bo) \<and> pair(**Lset(i),u,bo,x))]" |
|
923 |
apply (unfold is_bool_of_o_def) |
|
924 |
apply (intro FOL_reflections function_reflections Cons_reflection) |
|
925 |
done |
|
926 |
||
927 |
lemma Forall_replacement: |
|
928 |
"[|L(A); L(rp)|] |
|
929 |
==> strong_replacement |
|
930 |
(L, \<lambda>env z. \<exists>bo[L]. |
|
931 |
env \<in> list(A) & |
|
932 |
is_bool_of_o (L, |
|
933 |
\<forall>a[L]. \<forall>co[L]. \<forall>rpco[L]. |
|
934 |
a\<in>A --> is_Cons(L,a,env,co) --> |
|
935 |
fun_apply(L,rp,co,rpco) --> number1(L, rpco), |
|
936 |
bo) & |
|
937 |
pair(L,env,bo,z))" |
|
13566 | 938 |
apply (rule strong_replacementI) |
939 |
apply (rename_tac B) |
|
940 |
apply (rule_tac u="{A,list(A),B,rp}" in gen_separation [OF Forall_Reflects], |
|
941 |
simp add: list_closed) |
|
942 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13502 | 943 |
apply (rule DPow_LsetI) |
13566 | 944 |
apply (rule bex_iff_sats conj_iff_sats)+ |
945 |
apply (rule_tac env = "[u,x,A,list(A),B,rp]" in mem_iff_sats) |
|
946 |
apply (rule sep_rules is_bool_of_o_iff_sats Cons_iff_sats | simp)+ |
|
13502 | 947 |
done |
948 |
||
949 |
subsubsection{*The @{term "transrec_replacement"} Case*} |
|
950 |
||
13494 | 951 |
lemma formula_rec_replacement_Reflects: |
952 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L, u, y, x) \<and> |
|
953 |
is_wfrec (L, satisfies_MH(L,A), mesa, u, y)), |
|
954 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(**Lset(i), u, y, x) \<and> |
|
955 |
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
|
956 |
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
|
957 |
is_wfrec_reflection) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
958 |
|
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
959 |
lemma formula_rec_replacement: |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
960 |
--{*For the @{term transrec}*} |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
961 |
"[|n \<in> nat; L(A)|] ==> transrec_replacement(L, satisfies_MH(L,A), n)" |
13566 | 962 |
apply (rule transrec_replacementI, simp add: nat_into_M) |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
963 |
apply (rule strong_replacementI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
964 |
apply (rename_tac B) |
13566 | 965 |
apply (rule_tac u="{B,A,n,Memrel(eclose({n}))}" |
966 |
in gen_separation [OF formula_rec_replacement_Reflects], |
|
967 |
simp add: nat_into_M) |
|
968 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
969 |
apply (rule DPow_LsetI) |
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
970 |
apply (rule bex_iff_sats conj_iff_sats)+ |
13566 | 971 |
apply (rule_tac env = "[u,x,A,n,B,Memrel(eclose({n}))]" in mem_iff_sats) |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13494
diff
changeset
|
972 |
apply (rule sep_rules satisfies_MH_iff_sats is_wfrec_iff_sats | simp)+ |
13494 | 973 |
done |
974 |
||
13502 | 975 |
|
976 |
subsubsection{*The Lambda Replacement Case*} |
|
977 |
||
978 |
lemma formula_rec_lambda_replacement_Reflects: |
|
979 |
"REFLECTS [\<lambda>x. \<exists>u[L]. u \<in> B & |
|
980 |
mem_formula(L,u) & |
|
981 |
(\<exists>c[L]. |
|
982 |
is_formula_case |
|
983 |
(L, satisfies_is_a(L,A), satisfies_is_b(L,A), |
|
984 |
satisfies_is_c(L,A,g), satisfies_is_d(L,A,g), |
|
985 |
u, c) & |
|
986 |
pair(L,u,c,x)), |
|
987 |
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & mem_formula(**Lset(i),u) & |
|
988 |
(\<exists>c \<in> Lset(i). |
|
989 |
is_formula_case |
|
990 |
(**Lset(i), satisfies_is_a(**Lset(i),A), satisfies_is_b(**Lset(i),A), |
|
991 |
satisfies_is_c(**Lset(i),A,g), satisfies_is_d(**Lset(i),A,g), |
|
992 |
u, c) & |
|
993 |
pair(**Lset(i),u,c,x))]" |
|
994 |
by (intro FOL_reflections function_reflections mem_formula_reflection |
|
995 |
is_formula_case_reflection satisfies_is_a_reflection |
|
996 |
satisfies_is_b_reflection satisfies_is_c_reflection |
|
997 |
satisfies_is_d_reflection) |
|
998 |
||
999 |
lemma formula_rec_lambda_replacement: |
|
1000 |
--{*For the @{term transrec}*} |
|
1001 |
"[|L(g); L(A)|] ==> |
|
1002 |
strong_replacement (L, |
|
1003 |
\<lambda>x y. mem_formula(L,x) & |
|
1004 |
(\<exists>c[L]. is_formula_case(L, satisfies_is_a(L,A), |
|
1005 |
satisfies_is_b(L,A), |
|
1006 |
satisfies_is_c(L,A,g), |
|
1007 |
satisfies_is_d(L,A,g), x, c) & |
|
1008 |
pair(L, x, c, y)))" |
|
1009 |
apply (rule strong_replacementI) |
|
1010 |
apply (rename_tac B) |
|
13566 | 1011 |
apply (rule_tac u="{B,A,g}" |
1012 |
in gen_separation [OF formula_rec_lambda_replacement_Reflects], simp) |
|
1013 |
apply (drule mem_Lset_imp_subset_Lset, clarsimp) |
|
13502 | 1014 |
apply (rule DPow_LsetI) |
1015 |
apply (rule bex_iff_sats conj_iff_sats)+ |
|
13566 | 1016 |
apply (rule_tac env = "[u,x,A,g,B]" in mem_iff_sats) |
13502 | 1017 |
apply (rule sep_rules mem_formula_iff_sats |
1018 |
formula_case_iff_sats satisfies_is_a_iff_sats |
|
1019 |
satisfies_is_b_iff_sats satisfies_is_c_iff_sats |
|
1020 |
satisfies_is_d_iff_sats | simp)+ |
|
1021 |
done |
|
1022 |
||
1023 |
||
1024 |
subsection{*Instantiating @{text M_satisfies}*} |
|
1025 |
||
1026 |
lemma M_satisfies_axioms_L: "M_satisfies_axioms(L)" |
|
1027 |
apply (rule M_satisfies_axioms.intro) |
|
1028 |
apply (assumption | rule |
|
1029 |
Member_replacement Equal_replacement |
|
1030 |
Nand_replacement Forall_replacement |
|
1031 |
formula_rec_replacement formula_rec_lambda_replacement)+ |
|
1032 |
done |
|
1033 |
||
1034 |
theorem M_satisfies_L: "PROP M_satisfies(L)" |
|
1035 |
apply (rule M_satisfies.intro) |
|
1036 |
apply (rule M_eclose.axioms [OF M_eclose_L])+ |
|
1037 |
apply (rule M_satisfies_axioms_L) |
|
1038 |
done |
|
1039 |
||
13504 | 1040 |
text{*Finally: the point of the whole theory!*} |
1041 |
lemmas satisfies_closed = M_satisfies.satisfies_closed [OF M_satisfies_L] |
|
1042 |
and satisfies_abs = M_satisfies.satisfies_abs [OF M_satisfies_L] |
|
1043 |
||
13494 | 1044 |
end |