author | paulson |
Tue, 14 Oct 1997 17:23:01 +0200 | |
changeset 3862 | 6f389875ab33 |
parent 3731 | 71366483323b |
child 4091 | 771b1f6422a8 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/AC/DC.ML |
2 |
ID: $Id$ |
|
3 |
Author: Krzysztof Grabczewski |
|
1196 | 4 |
|
5 |
The proofs concerning the Axiom of Dependent Choice |
|
6 |
*) |
|
7 |
||
8 |
open DC; |
|
9 |
||
10 |
(* ********************************************************************** *) |
|
1461 | 11 |
(* DC ==> DC(omega) *) |
12 |
(* *) |
|
13 |
(* The scheme of the proof: *) |
|
14 |
(* *) |
|
15 |
(* Assume DC. Let R and x satisfy the premise of DC(omega). *) |
|
16 |
(* *) |
|
17 |
(* Define XX and RR as follows: *) |
|
18 |
(* *) |
|
19 |
(* XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) *) |
|
20 |
(* f RR g iff domain(g)=succ(domain(f)) & *) |
|
21 |
(* restrict(g, domain(f)) = f *) |
|
22 |
(* *) |
|
23 |
(* Then RR satisfies the hypotheses of DC. *) |
|
24 |
(* So applying DC: *) |
|
25 |
(* *) |
|
26 |
(* EX f:nat->XX. ALL n:nat. f`n RR f`succ(n) *) |
|
27 |
(* *) |
|
28 |
(* Thence *) |
|
29 |
(* *) |
|
30 |
(* ff = {<n, f`succ(n)`n>. n:nat} *) |
|
31 |
(* *) |
|
32 |
(* is the desired function. *) |
|
33 |
(* *) |
|
1196 | 34 |
(* ********************************************************************** *) |
35 |
||
2469 | 36 |
goal thy "{<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \ |
37 |
\ & restrict(z2, domain(z1)) = z1} <= XX*XX"; |
|
38 |
by (Fast_tac 1); |
|
1196 | 39 |
val lemma1_1 = result(); |
40 |
||
41 |
goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \ |
|
2469 | 42 |
\ ==> {<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \ |
1461 | 43 |
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \ |
2469 | 44 |
\ domain(z2)=succ(domain(z1)) \ |
45 |
\ & restrict(z2, domain(z1)) = z1} ~= 0"; |
|
1207 | 46 |
by (etac ballE 1); |
1196 | 47 |
by (eresolve_tac [empty_subsetI RS PowI RSN (2, notE)] 2); |
48 |
by (eresolve_tac [nat_0I RS n_lesspoll_nat RSN (2, impE)] 1); |
|
1207 | 49 |
by (etac bexE 1); |
1196 | 50 |
by (res_inst_tac [("a","<0, {<0, x>}>")] not_emptyI 1); |
1207 | 51 |
by (rtac CollectI 1); |
52 |
by (rtac SigmaI 1); |
|
2469 | 53 |
by (fast_tac (!claset addSIs [nat_0I RS UN_I, empty_fun]) 1); |
2493 | 54 |
by (rtac (nat_1I RS UN_I) 1); |
2469 | 55 |
by (fast_tac (!claset addSIs [singleton_fun RS Pi_type] |
56 |
addss (!simpset addsimps [singleton_0 RS sym])) 1); |
|
57 |
by (asm_full_simp_tac (!simpset addsimps [domain_0, domain_cons, |
|
58 |
singleton_0]) 1); |
|
1196 | 59 |
val lemma1_2 = result(); |
60 |
||
61 |
goal thy "!!X. ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \ |
|
2469 | 62 |
\ ==> range({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \ |
1461 | 63 |
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \ |
2469 | 64 |
\ domain(z2)=succ(domain(z1)) \ |
65 |
\ & restrict(z2, domain(z1)) = z1}) \ |
|
66 |
\ <= domain({<z1,z2>: (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}) * \ |
|
1461 | 67 |
\ (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}). \ |
2469 | 68 |
\ domain(z2)=succ(domain(z1)) \ |
69 |
\ & restrict(z2, domain(z1)) = z1})"; |
|
1207 | 70 |
by (rtac range_subset_domain 1); |
71 |
by (rtac subsetI 2); |
|
72 |
by (etac CollectD1 2); |
|
73 |
by (etac UN_E 1); |
|
74 |
by (etac CollectE 1); |
|
1196 | 75 |
by (dresolve_tac [fun_is_rel RS image_subset RS PowI RSN (2, bspec)] 1 |
1461 | 76 |
THEN (assume_tac 1)); |
1196 | 77 |
by (eresolve_tac [[n_lesspoll_nat, nat_into_Ord RSN (2, image_Ord_lepoll)] |
1461 | 78 |
MRS lepoll_lesspoll_lesspoll RSN (2, impE)] 1 |
79 |
THEN REPEAT (assume_tac 1)); |
|
1207 | 80 |
by (etac bexE 1); |
1196 | 81 |
by (res_inst_tac [("x","cons(<n,x>, g)")] exI 1); |
1207 | 82 |
by (rtac CollectI 1); |
83 |
by (rtac SigmaI 1); |
|
2469 | 84 |
by (Fast_tac 1); |
1207 | 85 |
by (rtac UN_I 1); |
86 |
by (etac nat_succI 1); |
|
87 |
by (rtac CollectI 1); |
|
88 |
by (etac cons_fun_type2 1 THEN (assume_tac 1)); |
|
2469 | 89 |
by (fast_tac (!claset addSEs [succE] addss (!simpset |
1461 | 90 |
addsimps [cons_image_n, cons_val_n, cons_image_k, cons_val_k])) 1); |
2469 | 91 |
by (asm_full_simp_tac (!simpset |
1461 | 92 |
addsimps [domain_cons, domain_of_fun, succ_def, restrict_cons_eq]) 1); |
1196 | 93 |
val lemma1_3 = result(); |
94 |
||
95 |
goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \ |
|
2469 | 96 |
\ RR = {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \ |
97 |
\ & restrict(z2, domain(z1)) = z1}; \ |
|
1461 | 98 |
\ ALL Y:Pow(X). Y lesspoll nat --> (EX x:X. <Y, x> : R) \ |
99 |
\ |] ==> RR <= XX*XX & RR ~= 0 & range(RR) <= domain(RR)"; |
|
2469 | 100 |
by (fast_tac (!claset addSIs [lemma1_1] addSEs [lemma1_2, lemma1_3]) 1); |
1196 | 101 |
val lemma1 = result(); |
102 |
||
2469 | 103 |
goal thy |
104 |
"!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \ |
|
1461 | 105 |
\ ALL n:nat. <f`n, f`succ(n)> : \ |
2469 | 106 |
\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \ |
107 |
\ & restrict(z2, domain(z1)) = z1}; \ |
|
1461 | 108 |
\ f: nat -> XX; n:nat \ |
2469 | 109 |
\ |] ==> EX k:nat. f`succ(n) : k -> X & n:k \ |
110 |
\ & <f`succ(n)``n, f`succ(n)`n> : R"; |
|
1207 | 111 |
by (etac nat_induct 1); |
1196 | 112 |
by (dresolve_tac [nat_1I RSN (2, apply_type)] 1); |
113 |
by (dresolve_tac [nat_0I RSN (2, bspec)] 1); |
|
2469 | 114 |
by (Asm_full_simp_tac 1); |
3731 | 115 |
by Safe_tac; |
1207 | 116 |
by (rtac bexI 1 THEN (assume_tac 2)); |
2469 | 117 |
by (best_tac (!claset addIs [ltD] |
2493 | 118 |
addSEs [nat_0_le RS leE] |
2469 | 119 |
addEs [sym RS trans RS succ_neq_0, domain_of_fun] |
2493 | 120 |
addss (!simpset)) 1); |
2469 | 121 |
(** LEVEL 7 **) |
1196 | 122 |
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); |
2469 | 123 |
by (subgoal_tac "f ` succ(succ(x)) : succ(k)->X" 1); |
1196 | 124 |
by (dresolve_tac [nat_succI RS nat_succI RSN (2, apply_type)] 1 |
1461 | 125 |
THEN (assume_tac 1)); |
2469 | 126 |
by (Full_simp_tac 1); |
3731 | 127 |
by Safe_tac; |
2469 | 128 |
by (forw_inst_tac [("a","succ(k)")] (domain_of_fun RS sym RS trans) 1 THEN |
129 |
(assume_tac 1)); |
|
130 |
by (forw_inst_tac [("a","xa")] (domain_of_fun RS sym RS trans) 1 THEN |
|
131 |
(assume_tac 1)); |
|
132 |
by (fast_tac (!claset addSEs [nat_succI, nat_into_Ord RS succ_in_succ] |
|
1461 | 133 |
addSDs [nat_into_Ord RS succ_in_succ RSN (2, bspec)]) 1); |
2469 | 134 |
by (dtac domain_of_fun 1); |
135 |
by (Full_simp_tac 1); |
|
136 |
by (deepen_tac (!claset addDs [domain_of_fun RS sym RS trans]) 0 1); |
|
1196 | 137 |
val lemma2 = result(); |
138 |
||
2469 | 139 |
goal thy |
140 |
"!!X.[| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \ |
|
1461 | 141 |
\ ALL n:nat. <f`n, f`succ(n)> : \ |
2469 | 142 |
\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \ |
143 |
\ & restrict(z2, domain(z1)) = z1}; \ |
|
1461 | 144 |
\ f: nat -> XX; n:nat \ |
2469 | 145 |
\ |] ==> ALL x:n. f`succ(n)`x = f`succ(x)`x"; |
1207 | 146 |
by (etac nat_induct 1); |
2469 | 147 |
by (Fast_tac 1); |
1207 | 148 |
by (rtac ballI 1); |
149 |
by (etac succE 1); |
|
150 |
by (rtac restrict_eq_imp_val_eq 1); |
|
1196 | 151 |
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); |
2469 | 152 |
by (Asm_full_simp_tac 1); |
1207 | 153 |
by (dtac lemma2 1 THEN REPEAT (assume_tac 1)); |
2469 | 154 |
by (fast_tac (!claset addSDs [domain_of_fun]) 1); |
1196 | 155 |
by (dres_inst_tac [("x","xa")] bspec 1 THEN (assume_tac 1)); |
156 |
by (eresolve_tac [sym RS trans RS sym] 1); |
|
157 |
by (resolve_tac [restrict_eq_imp_val_eq RS sym] 1); |
|
158 |
by (dresolve_tac [nat_succI RSN (2, bspec)] 1 THEN (assume_tac 1)); |
|
2469 | 159 |
by (Asm_full_simp_tac 1); |
1207 | 160 |
by (dtac lemma2 1 THEN REPEAT (assume_tac 1)); |
1196 | 161 |
by (fast_tac (FOL_cs addSDs [domain_of_fun] |
1461 | 162 |
addSEs [bexE, nat_into_Ord RSN (2, OrdmemD) RS subsetD]) 1); |
1196 | 163 |
val lemma3_1 = result(); |
164 |
||
165 |
goal thy "!!n. ALL x:n. f`succ(n)`x = f`succ(x)`x \ |
|
1461 | 166 |
\ ==> {f`succ(x)`x. x:n} = {f`succ(n)`x. x:n}"; |
2469 | 167 |
by (Asm_full_simp_tac 1); |
1196 | 168 |
val lemma3_2 = result(); |
169 |
||
170 |
goal thy "!!X. [| XX = (UN n:nat. {f:n->X. ALL k:n. <f``k, f`k> : R}); \ |
|
1461 | 171 |
\ ALL n:nat. <f`n, f`succ(n)> : \ |
2469 | 172 |
\ {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) \ |
173 |
\ & restrict(z2, domain(z1)) = z1}; \ |
|
1461 | 174 |
\ f: nat -> XX; n:nat \ |
175 |
\ |] ==> (lam x:nat. f`succ(x)`x) `` n = f`succ(n)``n"; |
|
1207 | 176 |
by (etac natE 1); |
2469 | 177 |
by (asm_full_simp_tac (!simpset addsimps [image_0]) 1); |
1196 | 178 |
by (resolve_tac [image_lam RS ssubst] 1); |
2469 | 179 |
by (fast_tac (!claset addSEs [[nat_succI, Ord_nat] MRS OrdmemD]) 1); |
1196 | 180 |
by (resolve_tac [lemma3_1 RS lemma3_2 RS ssubst] 1 |
1461 | 181 |
THEN REPEAT (assume_tac 1)); |
2469 | 182 |
by (fast_tac (!claset addSEs [nat_succI]) 1); |
1196 | 183 |
by (dresolve_tac [nat_succI RSN (4, lemma2)] 1 |
1461 | 184 |
THEN REPEAT (assume_tac 1)); |
2469 | 185 |
by (fast_tac (!claset addSEs [nat_into_Ord RSN (2, OrdmemD) RSN |
186 |
(2, image_fun RS sym)]) 1); |
|
1196 | 187 |
val lemma3 = result(); |
188 |
||
189 |
goal thy "!!f. [| f:A->B; B<=C |] ==> f:A->C"; |
|
1207 | 190 |
by (rtac Pi_type 1 THEN (assume_tac 1)); |
2469 | 191 |
by (fast_tac (!claset addSEs [apply_type]) 1); |
3731 | 192 |
qed "fun_type_gen"; |
1196 | 193 |
|
194 |
goalw thy [DC_def, DC0_def] "!!Z. DC0 ==> DC(nat)"; |
|
195 |
by (REPEAT (resolve_tac [allI, impI] 1)); |
|
196 |
by (REPEAT (eresolve_tac [conjE, allE] 1)); |
|
197 |
by (eresolve_tac [[refl, refl] MRS lemma1 RSN (2, impE)] 1 |
|
1461 | 198 |
THEN (assume_tac 1)); |
1207 | 199 |
by (etac bexE 1); |
1196 | 200 |
by (res_inst_tac [("x","lam n:nat. f`succ(n)`n")] bexI 1); |
2469 | 201 |
by (fast_tac (!claset addSIs [lam_type] addSDs [refl RS lemma2] |
1461 | 202 |
addSEs [fun_type_gen, apply_type]) 2); |
1207 | 203 |
by (rtac oallI 1); |
1196 | 204 |
by (forward_tac [ltD RSN (3, refl RS lemma2)] 1 |
1461 | 205 |
THEN assume_tac 2); |
2469 | 206 |
by (fast_tac (!claset addSEs [fun_type_gen]) 1); |
1196 | 207 |
by (eresolve_tac [ltD RSN (3, refl RS lemma3) RS ssubst] 1 |
1461 | 208 |
THEN assume_tac 2); |
2469 | 209 |
by (fast_tac (!claset addSEs [fun_type_gen]) 1); |
210 |
by (fast_tac (!claset addss (!simpset)) 1); |
|
1196 | 211 |
qed "DC0_DC_nat"; |
212 |
||
3862 | 213 |
(* ************************************************************************ |
214 |
DC(omega) ==> DC |
|
215 |
||
216 |
The scheme of the proof: |
|
217 |
||
218 |
Assume DC(omega). Let R and x satisfy the premise of DC. |
|
219 |
||
220 |
Define XX and RR as follows: |
|
221 |
||
222 |
XX = (UN n:nat. {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}) |
|
223 |
||
224 |
RR = {<z1,z2>:Fin(XX)*XX. |
|
225 |
(domain(z2)=succ(UN f:z1. domain(f)) & |
|
226 |
(ALL f:z1. restrict(z2, domain(f)) = f)) | |
|
227 |
(~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) & |
|
228 |
(ALL f:z1. restrict(g, domain(f)) = f)) & |
|
229 |
z2={<0,x>})} |
|
230 |
||
231 |
Then XX and RR satisfy the hypotheses of DC(omega). |
|
232 |
So applying DC: |
|
233 |
||
234 |
EX f:nat->XX. ALL n:nat. f``n RR f`n |
|
235 |
||
236 |
Thence |
|
237 |
||
238 |
ff = {<n, f`n`n>. n:nat} |
|
239 |
||
240 |
is the desired function. |
|
241 |
||
242 |
************************************************************************* *) |
|
1196 | 243 |
|
244 |
goalw thy [lesspoll_def, Finite_def] |
|
1461 | 245 |
"!!A. A lesspoll nat ==> Finite(A)"; |
2469 | 246 |
by (fast_tac (!claset addSDs [ltD, lepoll_imp_ex_le_eqpoll] |
1461 | 247 |
addSIs [Ord_nat]) 1); |
3731 | 248 |
qed "lesspoll_nat_is_Finite"; |
1196 | 249 |
|
250 |
goal thy "!!n. n:nat ==> ALL A. (A eqpoll n & A <= X) --> A : Fin(X)"; |
|
1207 | 251 |
by (etac nat_induct 1); |
252 |
by (rtac allI 1); |
|
2469 | 253 |
by (fast_tac (!claset addSIs [Fin.emptyI] |
1461 | 254 |
addSDs [eqpoll_imp_lepoll RS lepoll_0_is_0]) 1); |
1207 | 255 |
by (rtac allI 1); |
256 |
by (rtac impI 1); |
|
257 |
by (etac conjE 1); |
|
1196 | 258 |
by (resolve_tac [eqpoll_succ_imp_not_empty RS not_emptyE] 1 |
1461 | 259 |
THEN (assume_tac 1)); |
1196 | 260 |
by (forward_tac [Diff_sing_eqpoll] 1 THEN (assume_tac 1)); |
1207 | 261 |
by (etac allE 1); |
262 |
by (etac impE 1); |
|
2469 | 263 |
by (Fast_tac 1); |
1207 | 264 |
by (dtac subsetD 1 THEN (assume_tac 1)); |
1196 | 265 |
by (dresolve_tac [Fin.consI] 1 THEN (assume_tac 1)); |
2469 | 266 |
by (asm_full_simp_tac (!simpset addsimps [cons_Diff]) 1); |
3731 | 267 |
qed "Finite_Fin_lemma"; |
1196 | 268 |
|
269 |
goalw thy [Finite_def] "!!A. [| Finite(A); A <= X |] ==> A : Fin(X)"; |
|
1207 | 270 |
by (etac bexE 1); |
271 |
by (dtac Finite_Fin_lemma 1); |
|
272 |
by (etac allE 1); |
|
273 |
by (etac impE 1); |
|
1196 | 274 |
by (assume_tac 2); |
2469 | 275 |
by (Fast_tac 1); |
3731 | 276 |
qed "Finite_Fin"; |
1196 | 277 |
|
2469 | 278 |
goal thy "!!x. x: X \ |
279 |
\ ==> {<0,x>}: (UN n:nat. {f:succ(n)->X. ALL k:n. <f`k, f`succ(k)> : R})"; |
|
2493 | 280 |
by (rtac (nat_0I RS UN_I) 1); |
2469 | 281 |
by (fast_tac (!claset addSIs [singleton_fun RS Pi_type] |
282 |
addss (!simpset addsimps [singleton_0 RS sym])) 1); |
|
3731 | 283 |
qed "singleton_in_funs"; |
1196 | 284 |
|
285 |
goal thy |
|
1461 | 286 |
"!!X. [| XX = (UN n:nat. \ |
287 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
|
2469 | 288 |
\ RR = {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \ |
289 |
\ & (ALL f:z1. restrict(z2, domain(f)) = f)) | \ |
|
290 |
\ (~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) \ |
|
291 |
\ & (ALL f:z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}; \ |
|
1461 | 292 |
\ range(R) <= domain(R); x:domain(R) \ |
293 |
\ |] ==> RR <= Pow(XX)*XX & \ |
|
294 |
\ (ALL Y:Pow(XX). Y lesspoll nat --> (EX x:XX. <Y,x>:RR))"; |
|
1207 | 295 |
by (rtac conjI 1); |
2469 | 296 |
by (deepen_tac (!claset addSEs [FinD RS PowI]) 0 1); |
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
297 |
by (rtac (impI RS ballI) 1); |
1196 | 298 |
by (dresolve_tac [[lesspoll_nat_is_Finite, PowD] MRS Finite_Fin] 1 |
1461 | 299 |
THEN (assume_tac 1)); |
1196 | 300 |
by (excluded_middle_tac "EX g:XX. domain(g)=succ(UN f:Y. domain(f)) \ |
1461 | 301 |
\ & (ALL f:Y. restrict(g, domain(f)) = f)" 1); |
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
302 |
by (etac subst 2 THEN (*elimination equation for greater speed*) |
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
303 |
fast_tac (!claset addss (!simpset)) 2); |
3731 | 304 |
by (safe_tac (!claset delrules [domainE])); |
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
305 |
by (swap_res_tac [bexI] 1 THEN etac singleton_in_funs 2); |
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
306 |
by (asm_full_simp_tac (!simpset addsimps [nat_0I RSN (2, bexI), |
2493 | 307 |
cons_fun_type2, empty_fun]) 1); |
3862 | 308 |
val lemma4 = result(); |
1196 | 309 |
|
310 |
goal thy "!!f. [| f:nat->X; n:nat |] ==> \ |
|
1461 | 311 |
\ (UN x:f``succ(n). P(x)) = P(f`n) Un (UN x:f``n. P(x))"; |
2469 | 312 |
by (asm_full_simp_tac (!simpset |
1461 | 313 |
addsimps [Ord_nat RSN (2, OrdmemD) RSN (2, image_fun), |
314 |
[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); |
|
3731 | 315 |
qed "UN_image_succ_eq"; |
1196 | 316 |
|
317 |
goal thy "!!f. [| (UN x:f``n. P(x)) = y; P(f`n) = succ(y); \ |
|
1461 | 318 |
\ f:nat -> X; n:nat |] ==> (UN x:f``succ(n). P(x)) = succ(y)"; |
2469 | 319 |
by (asm_full_simp_tac (!simpset addsimps [UN_image_succ_eq]) 1); |
2496 | 320 |
by (Fast_tac 1); |
3731 | 321 |
qed "UN_image_succ_eq_succ"; |
1196 | 322 |
|
2469 | 323 |
goal thy "!!f. [| f:succ(n) -> D; n:nat; \ |
1461 | 324 |
\ domain(f)=succ(x); x=y |] ==> f`y : D"; |
2469 | 325 |
by (fast_tac (!claset addEs [apply_type] |
1461 | 326 |
addSDs [[sym, domain_of_fun] MRS trans]) 1); |
3731 | 327 |
qed "apply_domain_type"; |
1196 | 328 |
|
329 |
goal thy "!!f. [| f : nat -> X; n:nat |] ==> f``succ(n) = cons(f`n, f``n)"; |
|
2469 | 330 |
by (asm_full_simp_tac (!simpset |
1461 | 331 |
addsimps [nat_succI, Ord_nat RSN (2, OrdmemD), image_fun]) 1); |
3731 | 332 |
qed "image_fun_succ"; |
1196 | 333 |
|
334 |
goal thy "!!f. [| domain(f`n) = succ(u); f : nat -> (UN n:nat. \ |
|
1461 | 335 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
336 |
\ u=k; n:nat \ |
|
337 |
\ |] ==> f`n : succ(k) -> domain(R)"; |
|
1207 | 338 |
by (dtac apply_type 1 THEN (assume_tac 1)); |
2469 | 339 |
by (fast_tac (!claset addEs [UN_E, domain_eq_imp_fun_type]) 1); |
3731 | 340 |
qed "f_n_type"; |
1196 | 341 |
|
342 |
goal thy "!!f. [| f : nat -> (UN n:nat. \ |
|
1461 | 343 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
344 |
\ domain(f`n) = succ(k); n:nat \ |
|
345 |
\ |] ==> ALL i:k. <f`n`i, f`n`succ(i)> : R"; |
|
1207 | 346 |
by (dtac apply_type 1 THEN (assume_tac 1)); |
347 |
by (etac UN_E 1); |
|
348 |
by (etac CollectE 1); |
|
1196 | 349 |
by (dresolve_tac [domain_of_fun RS sym RS trans] 1 THEN (assume_tac 1)); |
2469 | 350 |
by (Asm_full_simp_tac 1); |
3731 | 351 |
qed "f_n_pairs_in_R"; |
1196 | 352 |
|
353 |
goalw thy [restrict_def] |
|
1461 | 354 |
"!!f. [| restrict(f, domain(x))=x; f:n->X; domain(x) <= n \ |
355 |
\ |] ==> restrict(cons(<n, y>, f), domain(x)) = x"; |
|
1196 | 356 |
by (eresolve_tac [sym RS trans RS sym] 1); |
1207 | 357 |
by (rtac fun_extension 1); |
2469 | 358 |
by (fast_tac (!claset addSIs [lam_type]) 1); |
359 |
by (fast_tac (!claset addSIs [lam_type]) 1); |
|
360 |
by (asm_full_simp_tac (!simpset addsimps [subsetD RS cons_val_k]) 1); |
|
3731 | 361 |
qed "restrict_cons_eq_restrict"; |
1196 | 362 |
|
363 |
goal thy "!!f. [| ALL x:f``n. restrict(f`n, domain(x))=x; \ |
|
1461 | 364 |
\ f : nat -> (UN n:nat. \ |
365 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
|
366 |
\ n:nat; domain(f`n) = succ(n); \ |
|
367 |
\ (UN x:f``n. domain(x)) <= n |] \ |
|
368 |
\ ==> ALL x:f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x))=x"; |
|
1207 | 369 |
by (rtac ballI 1); |
2469 | 370 |
by (asm_full_simp_tac (!simpset addsimps [image_fun_succ]) 1); |
1207 | 371 |
by (dtac f_n_type 1 THEN REPEAT (ares_tac [refl] 1)); |
2469 | 372 |
by (etac disjE 1); |
373 |
by (asm_full_simp_tac (!simpset addsimps [domain_of_fun, restrict_cons_eq]) 1); |
|
1207 | 374 |
by (dtac bspec 1 THEN (assume_tac 1)); |
2469 | 375 |
by (fast_tac (!claset addSEs [restrict_cons_eq_restrict]) 1); |
3731 | 376 |
qed "all_in_image_restrict_eq"; |
1196 | 377 |
|
2469 | 378 |
goal thy |
379 |
"!!X.[| ALL b<nat. <f``b, f`b> : \ |
|
3862 | 380 |
\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) & \ |
381 |
\ (ALL f:z1. restrict(z2, domain(f)) = f)) | \ |
|
382 |
\ (~ (EX g:XX. domain(g)=succ(UN f:z1. domain(f)) & \ |
|
383 |
\ (ALL f:z1. restrict(g, domain(f)) = f)) & \ |
|
384 |
\ z2={<0,x>})}; \ |
|
1461 | 385 |
\ XX = (UN n:nat. \ |
386 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
|
387 |
\ f: nat -> XX; range(R) <= domain(R); x:domain(R) \ |
|
2469 | 388 |
\ |] ==> ALL b<nat. <f``b, f`b> : \ |
389 |
\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \ |
|
390 |
\ & (UN f:z1. domain(f)) = b \ |
|
391 |
\ & (ALL f:z1. restrict(z2, domain(f)) = f))}"; |
|
1207 | 392 |
by (rtac oallI 1); |
393 |
by (dtac ltD 1); |
|
394 |
by (etac nat_induct 1); |
|
1196 | 395 |
by (dresolve_tac [[nat_0I, Ord_nat] MRS ltI RSN (2, ospec)] 1); |
1200 | 396 |
by (fast_tac (FOL_cs addss |
3862 | 397 |
(!simpset addsimps [singleton_fun RS domain_of_fun, |
2493 | 398 |
singleton_0, singleton_in_funs])) 1); |
2469 | 399 |
(*induction step*) (** LEVEL 5 **) |
3862 | 400 |
by (full_simp_tac (*prevent simplification of ~EX to ALL~*) |
401 |
(FOL_ss addsimps [separation, split]) 1); |
|
1196 | 402 |
by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 |
1461 | 403 |
THEN (assume_tac 1)); |
2469 | 404 |
by (REPEAT (eresolve_tac [conjE, disjE] 1)); |
1200 | 405 |
by (fast_tac (FOL_cs addSEs [trans, subst_context] |
2469 | 406 |
addSIs [UN_image_succ_eq_succ] addss (!simpset)) 1); |
1207 | 407 |
by (etac conjE 1); |
408 |
by (etac notE 1); |
|
2469 | 409 |
by (asm_full_simp_tac (!simpset addsimps [UN_image_succ_eq_succ]) 1); |
410 |
(** LEVEL 12 **) |
|
411 |
by (REPEAT (eresolve_tac [conjE, bexE] 1)); |
|
412 |
by (dtac apply_domain_type 1 THEN REPEAT (assume_tac 1)); |
|
1207 | 413 |
by (etac domainE 1); |
414 |
by (etac domainE 1); |
|
2469 | 415 |
|
1196 | 416 |
by (forward_tac [f_n_type] 1 THEN REPEAT (assume_tac 1)); |
3862 | 417 |
|
2493 | 418 |
by (rtac bexI 1); |
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
419 |
by (etac nat_succI 2); |
1196 | 420 |
by (res_inst_tac [("x","cons(<succ(xa), ya>, f`xa)")] bexI 1); |
2493 | 421 |
by (rtac conjI 1); |
1196 | 422 |
by (fast_tac (FOL_cs |
1461 | 423 |
addEs [subst_context RSN (2, trans) RS domain_cons_eq_succ, |
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
424 |
subst_context, all_in_image_restrict_eq, trans, equalityD1]) 2); |
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
425 |
by (eresolve_tac [rangeI RSN (2, subsetD) RSN (2, cons_fun_type2)] 2 |
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
426 |
THEN REPEAT (assume_tac 2)); |
1207 | 427 |
by (rtac ballI 1); |
428 |
by (etac succE 1); |
|
2469 | 429 |
(** LEVEL 25 **) |
430 |
by (dresolve_tac [domain_of_fun RSN (2, f_n_pairs_in_R)] 2 |
|
431 |
THEN REPEAT (assume_tac 2)); |
|
432 |
by (dtac bspec 2 THEN (assume_tac 2)); |
|
433 |
by (asm_full_simp_tac (!simpset |
|
434 |
addsimps [nat_into_Ord RS succ_in_succ, succI2, cons_val_k]) 2); |
|
435 |
by (asm_full_simp_tac (!simpset addsimps [cons_val_n, cons_val_k]) 1); |
|
3731 | 436 |
qed "simplify_recursion"; |
1196 | 437 |
|
2483
95c2f9c0930a
Default rewrite rules for quantification over Collect(A,P)
paulson
parents:
2469
diff
changeset
|
438 |
|
1196 | 439 |
goal thy "!!X. [| XX = (UN n:nat. \ |
1461 | 440 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
441 |
\ ALL b<nat. <f``b, f`b> : \ |
|
2469 | 442 |
\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \ |
443 |
\ & (UN f:z1. domain(f)) = b \ |
|
444 |
\ & (ALL f:z1. restrict(z2, domain(f)) = f))}; \ |
|
1461 | 445 |
\ f: nat -> XX; range(R) <= domain(R); x:domain(R); n:nat \ |
446 |
\ |] ==> f`n : succ(n) -> domain(R) \ |
|
447 |
\ & (ALL i:n. <f`n`i, f`n`succ(i)>:R)"; |
|
1207 | 448 |
by (dtac ospec 1); |
1196 | 449 |
by (eresolve_tac [Ord_nat RSN (2, ltI)] 1); |
1207 | 450 |
by (etac CollectE 1); |
2469 | 451 |
by (Asm_full_simp_tac 1); |
1207 | 452 |
by (rtac conjI 1); |
2469 | 453 |
by (fast_tac (!claset |
1461 | 454 |
addSEs [trans RS domain_eq_imp_fun_type, subst_context]) 1); |
2469 | 455 |
by (fast_tac (FOL_cs addSEs [conjE, f_n_pairs_in_R, trans, subst_context]) 1); |
1196 | 456 |
val lemma2 = result(); |
457 |
||
3862 | 458 |
goal thy |
459 |
"!!n. [| XX = (UN n:nat. \ |
|
460 |
\ {f:succ(n)->domain(R). ALL k:n. <f`k, f`succ(k)> : R}); \ |
|
1461 | 461 |
\ ALL b<nat. <f``b, f`b> : \ |
3862 | 462 |
\ {<z1,z2>:Fin(XX)*XX. (domain(z2)=succ(UN f:z1. domain(f)) \ |
2469 | 463 |
\ & (UN f:z1. domain(f)) = b \ |
464 |
\ & (ALL f:z1. restrict(z2, domain(f)) = f))}; \ |
|
3862 | 465 |
\ f : nat -> XX; n:nat; range(R) <= domain(R); x:domain(R) \ |
1461 | 466 |
\ |] ==> f`n`n = f`succ(n)`n"; |
1196 | 467 |
by (forward_tac [lemma2 RS conjunct1 RS domain_of_fun] 1 |
1461 | 468 |
THEN REPEAT (assume_tac 1)); |
1196 | 469 |
by (dresolve_tac [[nat_succI, Ord_nat] MRS ltI RSN (2, ospec)] 1 |
1461 | 470 |
THEN (assume_tac 1)); |
2469 | 471 |
by (Asm_full_simp_tac 1); |
1207 | 472 |
by (REPEAT (etac conjE 1)); |
473 |
by (etac ballE 1); |
|
1196 | 474 |
by (eresolve_tac [restrict_eq_imp_val_eq RS sym] 1); |
2469 | 475 |
by (fast_tac (!claset addSEs [ssubst]) 1); |
476 |
by (asm_full_simp_tac (!simpset |
|
1461 | 477 |
addsimps [[nat_succI, Ord_nat] MRS OrdmemD RSN (2, image_fun)]) 1); |
1196 | 478 |
val lemma3 = result(); |
479 |
||
3862 | 480 |
|
1196 | 481 |
goalw thy [DC_def, DC0_def] "!!Z. DC(nat) ==> DC0"; |
482 |
by (REPEAT (resolve_tac [allI, impI] 1)); |
|
483 |
by (REPEAT (eresolve_tac [asm_rl, conjE, ex_in_domain RS exE, allE] 1)); |
|
3862 | 484 |
by (eresolve_tac [[refl, refl] MRS lemma4 RSN (2, impE)] 1 |
1461 | 485 |
THEN REPEAT (assume_tac 1)); |
1207 | 486 |
by (etac bexE 1); |
1196 | 487 |
by (dresolve_tac [refl RSN (2, simplify_recursion)] 1 |
1461 | 488 |
THEN REPEAT (assume_tac 1)); |
1196 | 489 |
by (res_inst_tac [("x","lam n:nat. f`n`n")] bexI 1); |
1207 | 490 |
by (rtac lam_type 2); |
1196 | 491 |
by (eresolve_tac [[refl RS lemma2 RS conjunct1, succI1] MRS apply_type] 2 |
1461 | 492 |
THEN REPEAT (assume_tac 2)); |
1207 | 493 |
by (rtac ballI 1); |
1196 | 494 |
by (forward_tac [refl RS (nat_succI RSN (6, lemma2)) RS conjunct2] 1 |
1461 | 495 |
THEN REPEAT (assume_tac 1)); |
1196 | 496 |
by (dresolve_tac [refl RS lemma3] 1 THEN REPEAT (assume_tac 1)); |
2469 | 497 |
by (asm_full_simp_tac (!simpset addsimps [nat_succI]) 1); |
1196 | 498 |
qed "DC_nat_DC0"; |
499 |
||
500 |
(* ********************************************************************** *) |
|
1461 | 501 |
(* ALL K. Card(K) --> DC(K) ==> WO3 *) |
1196 | 502 |
(* ********************************************************************** *) |
503 |
||
504 |
goalw thy [lesspoll_def] |
|
1461 | 505 |
"!!A. [| ~ A lesspoll B; C lesspoll B |] ==> A - C ~= 0"; |
2469 | 506 |
by (fast_tac (!claset addSDs [Diff_eq_0_iff RS iffD1 RS subset_imp_lepoll] |
1461 | 507 |
addSIs [eqpollI] addEs [notE] addSEs [eqpollE, lepoll_trans]) 1); |
3862 | 508 |
val lesspoll_lemma = result(); |
1196 | 509 |
|
510 |
val [f_type, Ord_a, not_eq] = goalw thy [inj_def] |
|
1461 | 511 |
"[| f:a->X; Ord(a); (!!b c. [| b<c; c:a |] ==> f`b~=f`c) \ |
512 |
\ |] ==> f:inj(a, X)"; |
|
1196 | 513 |
by (resolve_tac [f_type RS CollectI] 1); |
514 |
by (REPEAT (resolve_tac [ballI,impI] 1)); |
|
515 |
by (resolve_tac [Ord_a RS Ord_in_Ord RS Ord_linear_lt] 1 |
|
1461 | 516 |
THEN (assume_tac 1)); |
1196 | 517 |
by (eres_inst_tac [("j","x")] (Ord_a RS Ord_in_Ord) 1); |
2469 | 518 |
by (REPEAT (fast_tac (!claset addDs [not_eq, not_eq RS not_sym]) 1)); |
3731 | 519 |
qed "fun_Ord_inj"; |
1196 | 520 |
|
521 |
goal thy "!!a. [| f:X->Y; A<=X; a:A |] ==> f`a : f``A"; |
|
2469 | 522 |
by (fast_tac (!claset addSEs [image_fun RS ssubst]) 1); |
3731 | 523 |
qed "value_in_image"; |
1196 | 524 |
|
525 |
goalw thy [DC_def, WO3_def] |
|
1461 | 526 |
"!!Z. ALL K. Card(K) --> DC(K) ==> WO3"; |
1207 | 527 |
by (rtac allI 1); |
1196 | 528 |
by (excluded_middle_tac "A lesspoll Hartog(A)" 1); |
2469 | 529 |
by (fast_tac (!claset addSDs [lesspoll_imp_ex_lt_eqpoll] |
1461 | 530 |
addSIs [Ord_Hartog, leI RS le_imp_subset]) 2); |
1196 | 531 |
by (REPEAT (eresolve_tac [allE, impE] 1)); |
1207 | 532 |
by (rtac Card_Hartog 1); |
1196 | 533 |
by (eres_inst_tac [("x","A")] allE 1); |
2469 | 534 |
by (eres_inst_tac [("x","{<z1,z2>:Pow(A)*A . z1 \ |
535 |
\ lesspoll Hartog(A) & z2 ~: z1}")] allE 1); |
|
536 |
by (Asm_full_simp_tac 1); |
|
1207 | 537 |
by (etac impE 1); |
3862 | 538 |
by (fast_tac (!claset addEs [lesspoll_lemma RS not_emptyE]) 1); |
1207 | 539 |
by (etac bexE 1); |
1196 | 540 |
by (resolve_tac [exI RS (lepoll_def RS (def_imp_iff RS iffD2)) |
1461 | 541 |
RS (HartogI RS notE)] 1); |
1196 | 542 |
by (resolve_tac [Ord_Hartog RSN (2, fun_Ord_inj)] 1 THEN (assume_tac 1)); |
543 |
by (dresolve_tac [Ord_Hartog RSN (2, OrdmemD) RSN (2, |
|
1461 | 544 |
ltD RSN (3, value_in_image))] 1 |
545 |
THEN REPEAT (assume_tac 1)); |
|
2469 | 546 |
by (fast_tac (!claset addSDs [Ord_Hartog RSN (2, ltI) RSN (2, ospec)] |
1461 | 547 |
addEs [subst]) 1); |
1196 | 548 |
qed "DC_WO3"; |
549 |
||
550 |
(* ********************************************************************** *) |
|
1461 | 551 |
(* WO1 ==> ALL K. Card(K) --> DC(K) *) |
1196 | 552 |
(* ********************************************************************** *) |
553 |
||
554 |
goal thy |
|
1461 | 555 |
"!!a. [| Ord(a); b:a |] ==> (lam x:a. P(x))``b = (lam x:b. P(x))``b"; |
1207 | 556 |
by (rtac images_eq 1); |
2469 | 557 |
by (REPEAT (fast_tac (!claset addSEs [RepFunI, OrdmemD] |
1461 | 558 |
addSIs [lam_type]) 2)); |
1207 | 559 |
by (rtac ballI 1); |
1196 | 560 |
by (dresolve_tac [OrdmemD RS subsetD] 1 |
1461 | 561 |
THEN REPEAT (assume_tac 1)); |
2469 | 562 |
by (Asm_full_simp_tac 1); |
3731 | 563 |
qed "lam_images_eq"; |
1196 | 564 |
|
565 |
goalw thy [lesspoll_def] "!!K. [| Card(K); b:K |] ==> b lesspoll K"; |
|
2469 | 566 |
by (asm_full_simp_tac (!simpset addsimps [Card_iff_initial]) 1); |
567 |
by (fast_tac (!claset addSIs [le_imp_lepoll, ltI, leI]) 1); |
|
3731 | 568 |
qed "in_Card_imp_lesspoll"; |
1196 | 569 |
|
570 |
goal thy "(lam b:a. P(b)) : a -> {P(b). b:a}"; |
|
2469 | 571 |
by (fast_tac (!claset addSIs [lam_type, RepFunI]) 1); |
3731 | 572 |
qed "lam_type_RepFun"; |
1196 | 573 |
|
574 |
goal thy "!!Z. [| ALL Y:Pow(X). Y lesspoll a --> (EX x:X. <Y, x> : R); \ |
|
1461 | 575 |
\ b:a; Z:Pow(X); Z lesspoll a |] \ |
576 |
\ ==> {x:X. <Z,x> : R} ~= 0"; |
|
1200 | 577 |
by (fast_tac (FOL_cs addEs [bexE, equals0D] |
1461 | 578 |
addSDs [bspec] addIs [CollectI]) 1); |
1196 | 579 |
val lemma_ = result(); |
580 |
||
581 |
goal thy "!!K. [| Card(K); well_ord(X,Q); \ |
|
1461 | 582 |
\ ALL Y:Pow(X). Y lesspoll K --> (EX x:X. <Y, x> : R); b:K |] \ |
583 |
\ ==> ff(b, X, Q, R) : {x:X. <(lam c:b. ff(c, X, Q, R))``b, x> : R}"; |
|
1196 | 584 |
by (res_inst_tac [("P","b:K")] impE 1 THEN TRYALL assume_tac); |
585 |
by (res_inst_tac [("i","b")] (Card_is_Ord RS Ord_in_Ord RS trans_induct) 1 |
|
1461 | 586 |
THEN REPEAT (assume_tac 1)); |
1207 | 587 |
by (rtac impI 1); |
1196 | 588 |
by (resolve_tac [ff_def RS def_transrec RS ssubst] 1); |
1207 | 589 |
by (etac the_first_in 1); |
2469 | 590 |
by (Fast_tac 1); |
591 |
by (asm_full_simp_tac (!simpset |
|
1461 | 592 |
addsimps [[lam_type_RepFun, subset_refl] MRS image_fun]) 1); |
1207 | 593 |
by (etac lemma_ 1 THEN (assume_tac 1)); |
2469 | 594 |
by (fast_tac (!claset addSEs [RepFunE, impE, notE] |
1461 | 595 |
addEs [Card_is_Ord RSN (2, OrdmemD) RS subsetD]) 1); |
2469 | 596 |
by (fast_tac (!claset addSEs [[in_Card_imp_lesspoll, RepFun_lepoll] |
1461 | 597 |
MRS lepoll_lesspoll_lesspoll]) 1); |
1196 | 598 |
val lemma = result(); |
599 |
||
600 |
goalw thy [DC_def, WO1_def] |
|
1461 | 601 |
"!!Z. WO1 ==> ALL K. Card(K) --> DC(K)"; |
1196 | 602 |
by (REPEAT (resolve_tac [allI,impI] 1)); |
603 |
by (REPEAT (eresolve_tac [allE,exE,conjE] 1)); |
|
604 |
by (res_inst_tac [("x","lam b:K. ff(b, X, Ra, R)")] bexI 1); |
|
1207 | 605 |
by (rtac lam_type 2); |
1196 | 606 |
by (resolve_tac [lemma RS CollectD1] 2 THEN REPEAT (assume_tac 2)); |
2469 | 607 |
by (asm_full_simp_tac (!simpset |
1461 | 608 |
addsimps [[Card_is_Ord, ltD] MRS lam_images_eq]) 1); |
2469 | 609 |
by (fast_tac (!claset addSEs [ltE, lemma RS CollectD2]) 1); |
1196 | 610 |
qed" WO1_DC_Card"; |
611 |