485
|
1 |
(* Title: ZF/Zorn.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
4 |
Copyright 1994 University of Cambridge
|
|
5 |
|
516
|
6 |
Proofs from the paper
|
485
|
7 |
Abrial & Laffitte,
|
|
8 |
Towards the Mechanization of the Proofs of Some
|
|
9 |
Classical Theorems of Set Theory.
|
|
10 |
*)
|
|
11 |
|
516
|
12 |
open Zorn;
|
485
|
13 |
|
516
|
14 |
(*** Section 1. Mathematical Preamble ***)
|
|
15 |
|
|
16 |
goal ZF.thy "!!A B C. (ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)";
|
|
17 |
by (fast_tac ZF_cs 1);
|
|
18 |
val Union_lemma0 = result();
|
|
19 |
|
|
20 |
goal ZF.thy
|
|
21 |
"!!A B C. [| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B";
|
|
22 |
by (fast_tac ZF_cs 1);
|
|
23 |
val Inter_lemma0 = result();
|
|
24 |
|
|
25 |
|
|
26 |
(*** Section 2. The Transfinite Construction ***)
|
|
27 |
|
|
28 |
goalw Zorn.thy [increasing_def]
|
|
29 |
"!!f A. f: increasing(A) ==> f: Pow(A)->Pow(A)";
|
|
30 |
by (eresolve_tac [CollectD1] 1);
|
|
31 |
val increasingD1 = result();
|
|
32 |
|
|
33 |
goalw Zorn.thy [increasing_def]
|
|
34 |
"!!f A. [| f: increasing(A); x<=A |] ==> x <= f`x";
|
|
35 |
by (eresolve_tac [CollectD2 RS spec RS mp] 1);
|
|
36 |
by (assume_tac 1);
|
|
37 |
val increasingD2 = result();
|
|
38 |
|
485
|
39 |
(*Introduction rules*)
|
516
|
40 |
val [TFin_nextI, Pow_TFin_UnionI] = TFin.intrs;
|
485
|
41 |
val TFin_UnionI = PowI RS Pow_TFin_UnionI;
|
|
42 |
|
516
|
43 |
val TFin_is_subset = TFin.dom_subset RS subsetD RS PowD;
|
485
|
44 |
|
|
45 |
|
|
46 |
(** Structural induction on TFin(S,next) **)
|
|
47 |
|
|
48 |
val major::prems = goal Zorn.thy
|
|
49 |
"[| n: TFin(S,next); \
|
|
50 |
\ !!x. [| x : TFin(S,next); P(x); next: increasing(S) |] ==> P(next`x); \
|
|
51 |
\ !!Y. [| Y <= TFin(S,next); ALL y:Y. P(y) |] ==> P(Union(Y)) \
|
|
52 |
\ |] ==> P(n)";
|
516
|
53 |
by (rtac (major RS TFin.induct) 1);
|
485
|
54 |
by (ALLGOALS (fast_tac (ZF_cs addIs prems)));
|
|
55 |
val TFin_induct = result();
|
|
56 |
|
|
57 |
(*Perform induction on n, then prove the major premise using prems. *)
|
|
58 |
fun TFin_ind_tac a prems i =
|
|
59 |
EVERY [res_inst_tac [("n",a)] TFin_induct i,
|
|
60 |
rename_last_tac a ["1"] (i+1),
|
|
61 |
rename_last_tac a ["2"] (i+2),
|
|
62 |
ares_tac prems i];
|
|
63 |
|
|
64 |
(*** Section 3. Some Properties of the Transfinite Construction ***)
|
|
65 |
|
|
66 |
val increasing_trans =
|
|
67 |
TFin_is_subset RSN (3, increasingD2 RSN (2,subset_trans)) |> standard;
|
|
68 |
|
|
69 |
(*Lemma 1 of section 3.1*)
|
|
70 |
val major::prems = goal Zorn.thy
|
|
71 |
"[| n: TFin(S,next); m: TFin(S,next); \
|
|
72 |
\ ALL x: TFin(S,next) . x<=m --> x=m | next`x<=m \
|
|
73 |
\ |] ==> n<=m | next`m<=n";
|
|
74 |
by (cut_facts_tac prems 1);
|
|
75 |
br (major RS TFin_induct) 1;
|
|
76 |
by (etac Union_lemma0 2); (*or just fast_tac ZF_cs*)
|
|
77 |
by (fast_tac (subset_cs addIs [increasing_trans]) 1);
|
|
78 |
val TFin_linear_lemma1 = result();
|
|
79 |
|
|
80 |
(*Lemma 2 of section 3.2. Interesting in its own right!
|
|
81 |
Requires next: increasing(S) in the second induction step. *)
|
|
82 |
val [major,ninc] = goal Zorn.thy
|
|
83 |
"[| m: TFin(S,next); next: increasing(S) \
|
|
84 |
\ |] ==> ALL n: TFin(S,next) . n<=m --> n=m | next`n<=m";
|
|
85 |
br (major RS TFin_induct) 1;
|
|
86 |
br (impI RS ballI) 1;
|
|
87 |
(*case split using TFin_linear_lemma1*)
|
|
88 |
by (res_inst_tac [("n1","n"), ("m1","x")]
|
|
89 |
(TFin_linear_lemma1 RS disjE) 1 THEN REPEAT (assume_tac 1));
|
|
90 |
by (dres_inst_tac [("x","n")] bspec 1 THEN assume_tac 1);
|
|
91 |
by (fast_tac (subset_cs addIs [increasing_trans]) 1);
|
|
92 |
by (REPEAT (ares_tac [disjI1,equalityI] 1));
|
|
93 |
(*second induction step*)
|
|
94 |
br (impI RS ballI) 1;
|
|
95 |
br (Union_lemma0 RS disjE) 1;
|
|
96 |
be disjI2 3;
|
|
97 |
by (REPEAT (ares_tac [disjI1,equalityI] 2));
|
|
98 |
br ballI 1;
|
|
99 |
by (ball_tac 1);
|
|
100 |
by (set_mp_tac 1);
|
|
101 |
by (res_inst_tac [("n1","n"), ("m1","x")]
|
|
102 |
(TFin_linear_lemma1 RS disjE) 1 THEN REPEAT (assume_tac 1));
|
|
103 |
by (fast_tac subset_cs 1);
|
|
104 |
br (ninc RS increasingD2 RS subset_trans RS disjI1) 1;
|
|
105 |
by (REPEAT (ares_tac [TFin_is_subset] 1));
|
|
106 |
val TFin_linear_lemma2 = result();
|
|
107 |
|
|
108 |
(*a more convenient form for Lemma 2*)
|
|
109 |
goal Zorn.thy
|
|
110 |
"!!m n. [| n<=m; m: TFin(S,next); n: TFin(S,next); next: increasing(S) \
|
|
111 |
\ |] ==> n=m | next`n<=m";
|
|
112 |
br (TFin_linear_lemma2 RS bspec RS mp) 1;
|
|
113 |
by (REPEAT (assume_tac 1));
|
|
114 |
val TFin_subsetD = result();
|
|
115 |
|
|
116 |
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*)
|
|
117 |
goal Zorn.thy
|
|
118 |
"!!m n. [| m: TFin(S,next); n: TFin(S,next); next: increasing(S) \
|
|
119 |
\ |] ==> n<=m | m<=n";
|
|
120 |
br (TFin_linear_lemma2 RSN (3,TFin_linear_lemma1) RS disjE) 1;
|
|
121 |
by (REPEAT (assume_tac 1) THEN etac disjI2 1);
|
|
122 |
by (fast_tac (subset_cs addIs [increasingD2 RS subset_trans,
|
|
123 |
TFin_is_subset]) 1);
|
|
124 |
val TFin_subset_linear = result();
|
|
125 |
|
|
126 |
|
|
127 |
(*Lemma 3 of section 3.3*)
|
|
128 |
val major::prems = goal Zorn.thy
|
|
129 |
"[| n: TFin(S,next); m: TFin(S,next); m = next`m |] ==> n<=m";
|
|
130 |
by (cut_facts_tac prems 1);
|
|
131 |
br (major RS TFin_induct) 1;
|
|
132 |
bd TFin_subsetD 1;
|
|
133 |
by (REPEAT (assume_tac 1));
|
|
134 |
by (fast_tac (ZF_cs addEs [ssubst]) 1);
|
|
135 |
by (fast_tac (subset_cs addIs [TFin_is_subset]) 1);
|
|
136 |
val equal_next_upper = result();
|
|
137 |
|
|
138 |
(*Property 3.3 of section 3.3*)
|
|
139 |
goal Zorn.thy
|
|
140 |
"!!m. [| m: TFin(S,next); next: increasing(S) \
|
|
141 |
\ |] ==> m = next`m <-> m = Union(TFin(S,next))";
|
|
142 |
br iffI 1;
|
|
143 |
br (Union_upper RS equalityI) 1;
|
|
144 |
br (equal_next_upper RS Union_least) 2;
|
|
145 |
by (REPEAT (assume_tac 1));
|
|
146 |
be ssubst 1;
|
|
147 |
by (rtac (increasingD2 RS equalityI) 1 THEN assume_tac 1);
|
|
148 |
by (ALLGOALS
|
|
149 |
(fast_tac (subset_cs addIs [TFin_UnionI, TFin_nextI, TFin_is_subset])));
|
|
150 |
val equal_next_Union = result();
|
|
151 |
|
|
152 |
|
|
153 |
(*** Section 4. Hausdorff's Theorem: every set contains a maximal chain ***)
|
|
154 |
(*** NB: We assume the partial ordering is <=, the subset relation! **)
|
|
155 |
|
|
156 |
(** Defining the "next" operation for Hausdorff's Theorem **)
|
|
157 |
|
|
158 |
goalw Zorn.thy [chain_def] "chain(A) <= Pow(A)";
|
|
159 |
by (resolve_tac [Collect_subset] 1);
|
|
160 |
val chain_subset_Pow = result();
|
|
161 |
|
|
162 |
goalw Zorn.thy [super_def] "super(A,c) <= chain(A)";
|
|
163 |
by (resolve_tac [Collect_subset] 1);
|
|
164 |
val super_subset_chain = result();
|
|
165 |
|
|
166 |
goalw Zorn.thy [maxchain_def] "maxchain(A) <= chain(A)";
|
|
167 |
by (resolve_tac [Collect_subset] 1);
|
|
168 |
val maxchain_subset_chain = result();
|
|
169 |
|
|
170 |
goal Zorn.thy
|
|
171 |
"!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \
|
|
172 |
\ X : chain(S); X ~: maxchain(S) \
|
|
173 |
\ |] ==> ch ` super(S,X) : super(S,X)";
|
|
174 |
by (eresolve_tac [apply_type] 1);
|
|
175 |
by (rewrite_goals_tac [super_def, maxchain_def]);
|
|
176 |
by (fast_tac ZF_cs 1);
|
|
177 |
val choice_super = result();
|
|
178 |
|
|
179 |
goal Zorn.thy
|
|
180 |
"!!S. [| ch : (PROD X:Pow(chain(S)) - {0}. X); \
|
|
181 |
\ X : chain(S); X ~: maxchain(S) \
|
|
182 |
\ |] ==> ch ` super(S,X) ~= X";
|
|
183 |
by (resolve_tac [notI] 1);
|
|
184 |
by (dresolve_tac [choice_super] 1);
|
|
185 |
by (assume_tac 1);
|
|
186 |
by (assume_tac 1);
|
|
187 |
by (asm_full_simp_tac (ZF_ss addsimps [super_def]) 1);
|
|
188 |
val choice_not_equals = result();
|
|
189 |
|
|
190 |
(*This justifies Definition 4.4*)
|
|
191 |
goal Zorn.thy
|
|
192 |
"!!S. ch: (PROD X: Pow(chain(S))-{0}. X) ==> \
|
|
193 |
\ EX next: increasing(S). ALL X: Pow(S). \
|
|
194 |
\ next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)";
|
|
195 |
by (rtac bexI 1);
|
|
196 |
by (rtac ballI 1);
|
|
197 |
by (resolve_tac [beta] 1);
|
|
198 |
by (assume_tac 1);
|
|
199 |
bw increasing_def;
|
|
200 |
by (rtac CollectI 1);
|
|
201 |
by (rtac lam_type 1);
|
|
202 |
by (asm_simp_tac (ZF_ss setloop split_tac [expand_if]) 1);
|
|
203 |
by (fast_tac (ZF_cs addSIs [super_subset_chain RS subsetD,
|
|
204 |
chain_subset_Pow RS subsetD,
|
|
205 |
choice_super]) 1);
|
|
206 |
(*Now, verify that it increases*)
|
|
207 |
by (resolve_tac [allI] 1);
|
|
208 |
by (resolve_tac [impI] 1);
|
|
209 |
by (asm_simp_tac (ZF_ss addsimps [Pow_iff, subset_refl]
|
|
210 |
setloop split_tac [expand_if]) 1);
|
|
211 |
by (safe_tac ZF_cs);
|
|
212 |
by (dresolve_tac [choice_super] 1);
|
|
213 |
by (REPEAT (assume_tac 1));
|
|
214 |
bw super_def;
|
|
215 |
by (fast_tac ZF_cs 1);
|
|
216 |
val Hausdorff_next_exists = result();
|
|
217 |
|
|
218 |
(*Lemma 4*)
|
|
219 |
goal Zorn.thy
|
|
220 |
"!!S. [| c: TFin(S,next); \
|
|
221 |
\ ch: (PROD X: Pow(chain(S))-{0}. X); \
|
|
222 |
\ next: increasing(S); \
|
|
223 |
\ ALL X: Pow(S). next`X = \
|
|
224 |
\ if(X: chain(S)-maxchain(S), ch`super(S,X), X) \
|
|
225 |
\ |] ==> c: chain(S)";
|
|
226 |
by (eresolve_tac [TFin_induct] 1);
|
|
227 |
by (asm_simp_tac
|
|
228 |
(ZF_ss addsimps [chain_subset_Pow RS subsetD,
|
|
229 |
choice_super RS (super_subset_chain RS subsetD)]
|
|
230 |
setloop split_tac [expand_if]) 1);
|
|
231 |
bw chain_def;
|
|
232 |
by (rtac CollectI 1 THEN fast_tac ZF_cs 1);
|
|
233 |
(*Cannot use safe_tac: the disjunction must be left alone*)
|
|
234 |
by (REPEAT (rtac ballI 1 ORELSE etac UnionE 1));
|
|
235 |
by (res_inst_tac [("m1","B"), ("n1","Ba")] (TFin_subset_linear RS disjE) 1);
|
|
236 |
(*fast_tac is just too slow here!*)
|
|
237 |
by (DEPTH_SOLVE (eresolve_tac [asm_rl, subsetD] 1
|
|
238 |
ORELSE ball_tac 1 THEN etac (CollectD2 RS bspec RS bspec) 1));
|
|
239 |
val TFin_chain_lemma4 = result();
|
|
240 |
|
|
241 |
goal Zorn.thy "EX c. c : maxchain(S)";
|
|
242 |
by (rtac (AC_Pi_Pow RS exE) 1);
|
|
243 |
by (rtac (Hausdorff_next_exists RS bexE) 1);
|
|
244 |
by (assume_tac 1);
|
|
245 |
by (rename_tac "ch next" 1);
|
|
246 |
by (subgoal_tac "Union(TFin(S,next)) : chain(S)" 1);
|
|
247 |
by (REPEAT (ares_tac [TFin_chain_lemma4, subset_refl RS TFin_UnionI] 2));
|
|
248 |
by (res_inst_tac [("x", "Union(TFin(S,next))")] exI 1);
|
|
249 |
by (resolve_tac [classical] 1);
|
|
250 |
by (subgoal_tac "next ` Union(TFin(S,next)) = Union(TFin(S,next))" 1);
|
|
251 |
by (resolve_tac [equal_next_Union RS iffD2 RS sym] 2);
|
|
252 |
by (resolve_tac [subset_refl RS TFin_UnionI] 2);
|
|
253 |
by (assume_tac 2);
|
|
254 |
by (resolve_tac [refl] 2);
|
|
255 |
by (asm_full_simp_tac
|
|
256 |
(ZF_ss addsimps [subset_refl RS TFin_UnionI RS
|
516
|
257 |
(TFin.dom_subset RS subsetD)]
|
485
|
258 |
setloop split_tac [expand_if]) 1);
|
|
259 |
by (eresolve_tac [choice_not_equals RS notE] 1);
|
|
260 |
by (REPEAT (assume_tac 1));
|
|
261 |
val Hausdorff = result();
|
|
262 |
|
|
263 |
|
|
264 |
(*** Section 5. Zorn's Lemma: if all chains in S have upper bounds in S
|
|
265 |
then S contains a maximal element ***)
|
|
266 |
|
|
267 |
(*Used in the proof of Zorn's Lemma*)
|
|
268 |
goalw Zorn.thy [chain_def]
|
|
269 |
"!!c. [| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)";
|
|
270 |
by (fast_tac ZF_cs 1);
|
|
271 |
val chain_extend = result();
|
|
272 |
|
|
273 |
goal Zorn.thy
|
|
274 |
"!!S. ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z";
|
|
275 |
by (resolve_tac [Hausdorff RS exE] 1);
|
|
276 |
by (asm_full_simp_tac (ZF_ss addsimps [maxchain_def]) 1);
|
|
277 |
by (rename_tac "c" 1);
|
|
278 |
by (res_inst_tac [("x", "Union(c)")] bexI 1);
|
|
279 |
by (fast_tac ZF_cs 2);
|
|
280 |
by (safe_tac ZF_cs);
|
|
281 |
by (rename_tac "z" 1);
|
|
282 |
by (resolve_tac [classical] 1);
|
|
283 |
by (subgoal_tac "cons(z,c): super(S,c)" 1);
|
|
284 |
by (fast_tac (ZF_cs addEs [equalityE]) 1);
|
|
285 |
bw super_def;
|
|
286 |
by (safe_tac eq_cs);
|
|
287 |
by (fast_tac (ZF_cs addEs [chain_extend]) 1);
|
|
288 |
by (best_tac (ZF_cs addEs [equalityE]) 1);
|
|
289 |
val Zorn = result();
|
|
290 |
|
|
291 |
|
|
292 |
(*** Section 6. Zermelo's Theorem: every set can be well-ordered ***)
|
|
293 |
|
|
294 |
(*Lemma 5*)
|
|
295 |
val major::prems = goal Zorn.thy
|
|
296 |
"[| n: TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) : Z \
|
|
297 |
\ |] ==> ALL m:Z. n<=m";
|
|
298 |
by (cut_facts_tac prems 1);
|
|
299 |
br (major RS TFin_induct) 1;
|
|
300 |
by (fast_tac ZF_cs 2); (*second induction step is easy*)
|
|
301 |
br ballI 1;
|
|
302 |
br (bspec RS TFin_subsetD RS disjE) 1;
|
|
303 |
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD]));
|
|
304 |
by (subgoal_tac "x = Inter(Z)" 1);
|
|
305 |
by (fast_tac ZF_cs 1);
|
|
306 |
by (fast_tac eq_cs 1);
|
|
307 |
val TFin_well_lemma5 = result();
|
|
308 |
|
|
309 |
(*Well-ordering of TFin(S,next)*)
|
|
310 |
goal Zorn.thy "!!Z. [| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z";
|
|
311 |
br classical 1;
|
|
312 |
by (subgoal_tac "Z = {Union(TFin(S,next))}" 1);
|
|
313 |
by (asm_simp_tac (ZF_ss addsimps [Inter_singleton]) 1);
|
|
314 |
be equal_singleton 1;
|
|
315 |
br (Union_upper RS equalityI) 1;
|
|
316 |
br (subset_refl RS TFin_UnionI RS TFin_well_lemma5 RS bspec) 2;
|
|
317 |
by (REPEAT_SOME (eresolve_tac [asm_rl,subsetD]));
|
|
318 |
val well_ord_TFin_lemma = result();
|
|
319 |
|
|
320 |
(*This theorem just packages the previous result*)
|
|
321 |
goal Zorn.thy
|
|
322 |
"!!S. next: increasing(S) ==> \
|
|
323 |
\ well_ord(TFin(S,next), Subset_rel(TFin(S,next)))";
|
|
324 |
by (resolve_tac [well_ordI] 1);
|
|
325 |
by (rewrite_goals_tac [Subset_rel_def, linear_def]);
|
|
326 |
(*Prove the linearity goal first*)
|
|
327 |
by (REPEAT (rtac ballI 2));
|
|
328 |
by (excluded_middle_tac "x=y" 2);
|
|
329 |
by (fast_tac ZF_cs 3);
|
|
330 |
(*The x~=y case remains*)
|
|
331 |
by (res_inst_tac [("n1","x"), ("m1","y")]
|
|
332 |
(TFin_subset_linear RS disjE) 2 THEN REPEAT (assume_tac 2));
|
|
333 |
by (fast_tac ZF_cs 2);
|
|
334 |
by (fast_tac ZF_cs 2);
|
|
335 |
(*Now prove the well_foundedness goal*)
|
|
336 |
by (resolve_tac [wf_onI] 1);
|
|
337 |
by (forward_tac [well_ord_TFin_lemma] 1 THEN assume_tac 1);
|
|
338 |
by (dres_inst_tac [("x","Inter(Z)")] bspec 1 THEN assume_tac 1);
|
|
339 |
by (fast_tac eq_cs 1);
|
|
340 |
val well_ord_TFin = result();
|
|
341 |
|
|
342 |
(** Defining the "next" operation for Zermelo's Theorem **)
|
|
343 |
|
|
344 |
goal AC.thy
|
|
345 |
"!!S. [| ch : (PROD X:Pow(S) - {0}. X); X<=S; X~=S \
|
|
346 |
\ |] ==> ch ` (S-X) : S-X";
|
|
347 |
by (eresolve_tac [apply_type] 1);
|
|
348 |
by (fast_tac (eq_cs addEs [equalityE]) 1);
|
|
349 |
val choice_Diff = result();
|
|
350 |
|
|
351 |
(*This justifies Definition 6.1*)
|
|
352 |
goal Zorn.thy
|
|
353 |
"!!S. ch: (PROD X: Pow(S)-{0}. X) ==> \
|
|
354 |
\ EX next: increasing(S). ALL X: Pow(S). \
|
|
355 |
\ next`X = if(X=S, S, cons(ch`(S-X), X))";
|
|
356 |
by (rtac bexI 1);
|
|
357 |
by (rtac ballI 1);
|
|
358 |
by (resolve_tac [beta] 1);
|
|
359 |
by (assume_tac 1);
|
|
360 |
bw increasing_def;
|
|
361 |
by (rtac CollectI 1);
|
|
362 |
by (rtac lam_type 1);
|
|
363 |
(*Verify that it increases*)
|
|
364 |
by (resolve_tac [allI] 2);
|
|
365 |
by (resolve_tac [impI] 2);
|
|
366 |
by (asm_simp_tac (ZF_ss addsimps [Pow_iff, subset_consI, subset_refl]
|
|
367 |
setloop split_tac [expand_if]) 2);
|
|
368 |
(*Type checking is surprisingly hard!*)
|
|
369 |
by (asm_simp_tac (ZF_ss addsimps [Pow_iff, cons_subset_iff, subset_refl]
|
|
370 |
setloop split_tac [expand_if]) 1);
|
|
371 |
by (fast_tac (ZF_cs addSIs [choice_Diff RS DiffD1]) 1);
|
|
372 |
val Zermelo_next_exists = result();
|
|
373 |
|
|
374 |
|
|
375 |
(*The construction of the injection*)
|
|
376 |
goal Zorn.thy
|
|
377 |
"!!S. [| ch: (PROD X: Pow(S)-{0}. X); \
|
|
378 |
\ next: increasing(S); \
|
|
379 |
\ ALL X: Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X)) \
|
|
380 |
\ |] ==> (lam x:S. Union({y: TFin(S,next). x~: y})) \
|
|
381 |
\ : inj(S, TFin(S,next) - {S})";
|
|
382 |
by (res_inst_tac [("d", "%y. ch`(S-y)")] lam_injective 1);
|
|
383 |
by (rtac DiffI 1);
|
|
384 |
by (resolve_tac [Collect_subset RS TFin_UnionI] 1);
|
|
385 |
by (fast_tac (ZF_cs addSIs [Collect_subset RS TFin_UnionI]
|
|
386 |
addEs [equalityE]) 1);
|
|
387 |
by (subgoal_tac "x ~: Union({y: TFin(S,next). x~: y})" 1);
|
|
388 |
by (fast_tac (ZF_cs addEs [equalityE]) 2);
|
|
389 |
by (subgoal_tac "Union({y: TFin(S,next). x~: y}) ~= S" 1);
|
|
390 |
by (fast_tac (ZF_cs addEs [equalityE]) 2);
|
|
391 |
(*For proving x : next`Union(...);
|
|
392 |
Abrial & Laffitte's justification appears to be faulty.*)
|
|
393 |
by (subgoal_tac "~ next ` Union({y: TFin(S,next). x~: y}) <= \
|
|
394 |
\ Union({y: TFin(S,next). x~: y})" 1);
|
|
395 |
by (asm_simp_tac
|
|
396 |
(ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
|
|
397 |
Pow_iff, cons_subset_iff, subset_refl,
|
|
398 |
choice_Diff RS DiffD2]
|
|
399 |
setloop split_tac [expand_if]) 2);
|
|
400 |
by (subgoal_tac "x : next ` Union({y: TFin(S,next). x~: y})" 1);
|
|
401 |
by (fast_tac (subset_cs addSIs [Collect_subset RS TFin_UnionI, TFin_nextI]) 2);
|
|
402 |
(*End of the lemmas!*)
|
|
403 |
by (asm_full_simp_tac
|
|
404 |
(ZF_ss addsimps [Collect_subset RS TFin_UnionI RS TFin_is_subset,
|
|
405 |
Pow_iff, cons_subset_iff, subset_refl]
|
|
406 |
setloop split_tac [expand_if]) 1);
|
|
407 |
by (REPEAT (eresolve_tac [asm_rl, consE, sym, notE] 1));
|
|
408 |
val choice_imp_injection = result();
|
|
409 |
|
|
410 |
(*The wellordering theorem*)
|
|
411 |
goal Zorn.thy "EX r. well_ord(S,r)";
|
|
412 |
by (rtac (AC_Pi_Pow RS exE) 1);
|
|
413 |
by (rtac (Zermelo_next_exists RS bexE) 1);
|
|
414 |
by (assume_tac 1);
|
|
415 |
br exI 1;
|
|
416 |
by (resolve_tac [well_ord_rvimage] 1);
|
|
417 |
by (eresolve_tac [well_ord_TFin] 2);
|
|
418 |
by (resolve_tac [choice_imp_injection RS inj_weaken_type] 1);
|
|
419 |
by (REPEAT (ares_tac [Diff_subset] 1));
|
|
420 |
val AC_well_ord = result();
|