author | paulson |
Tue, 13 Aug 2002 11:03:11 +0200 | |
changeset 13494 | 1c44289716ae |
parent 13356 | c9cfe1638bf2 |
child 13558 | 18acbbd4d225 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/Zorn.thy |
516 | 2 |
ID: $Id$ |
1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
516 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
*) |
|
7 |
||
13356 | 8 |
header{*Zorn's Lemma*} |
9 |
||
13134 | 10 |
theory Zorn = OrderArith + AC + Inductive: |
516 | 11 |
|
13356 | 12 |
text{*Based upon the unpublished article ``Towards the Mechanization of the |
13 |
Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.*} |
|
14 |
||
13134 | 15 |
constdefs |
16 |
Subset_rel :: "i=>i" |
|
17 |
"Subset_rel(A) == {z: A*A . EX x y. z=<x,y> & x<=y & x~=y}" |
|
18 |
||
19 |
chain :: "i=>i" |
|
20 |
"chain(A) == {F: Pow(A). ALL X:F. ALL Y:F. X<=Y | Y<=X}" |
|
516 | 21 |
|
13134 | 22 |
maxchain :: "i=>i" |
23 |
"maxchain(A) == {c: chain(A). super(A,c)=0}" |
|
24 |
||
25 |
super :: "[i,i]=>i" |
|
26 |
"super(A,c) == {d: chain(A). c<=d & c~=d}" |
|
485 | 27 |
|
516 | 28 |
|
13134 | 29 |
constdefs |
30 |
increasing :: "i=>i" |
|
31 |
"increasing(A) == {f: Pow(A)->Pow(A). ALL x. x<=A --> x<=f`x}" |
|
516 | 32 |
|
13356 | 33 |
|
34 |
(*Lemma for the inductive definition below*) |
|
35 |
lemma Union_in_Pow: "Y : Pow(Pow(A)) ==> Union(Y) : Pow(A)" |
|
36 |
by blast |
|
37 |
||
38 |
||
39 |
text{*We could make the inductive definition conditional on |
|
40 |
@{term "next \<in> increasing(S)"} |
|
516 | 41 |
but instead we make this a side-condition of an introduction rule. Thus |
42 |
the induction rule lets us assume that condition! Many inductive proofs |
|
13356 | 43 |
are therefore unconditional.*} |
516 | 44 |
consts |
13134 | 45 |
"TFin" :: "[i,i]=>i" |
516 | 46 |
|
47 |
inductive |
|
48 |
domains "TFin(S,next)" <= "Pow(S)" |
|
13134 | 49 |
intros |
50 |
nextI: "[| x : TFin(S,next); next: increasing(S) |] |
|
51 |
==> next`x : TFin(S,next)" |
|
516 | 52 |
|
13134 | 53 |
Pow_UnionI: "Y : Pow(TFin(S,next)) ==> Union(Y) : TFin(S,next)" |
516 | 54 |
|
6053
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
1478
diff
changeset
|
55 |
monos Pow_mono |
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
paulson
parents:
1478
diff
changeset
|
56 |
con_defs increasing_def |
13134 | 57 |
type_intros CollectD1 [THEN apply_funtype] Union_in_Pow |
58 |
||
59 |
||
13356 | 60 |
subsection{*Mathematical Preamble *} |
13134 | 61 |
|
62 |
lemma Union_lemma0: "(ALL x:C. x<=A | B<=x) ==> Union(C)<=A | B<=Union(C)" |
|
13269 | 63 |
by blast |
13134 | 64 |
|
13356 | 65 |
lemma Inter_lemma0: |
66 |
"[| c:C; ALL x:C. A<=x | x<=B |] ==> A<=Inter(C) | Inter(C)<=B" |
|
13269 | 67 |
by blast |
13134 | 68 |
|
69 |
||
13356 | 70 |
subsection{*The Transfinite Construction *} |
13134 | 71 |
|
72 |
lemma increasingD1: "f: increasing(A) ==> f: Pow(A)->Pow(A)" |
|
73 |
apply (unfold increasing_def) |
|
74 |
apply (erule CollectD1) |
|
75 |
done |
|
76 |
||
77 |
lemma increasingD2: "[| f: increasing(A); x<=A |] ==> x <= f`x" |
|
13269 | 78 |
by (unfold increasing_def, blast) |
13134 | 79 |
|
80 |
lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI, standard] |
|
81 |
||
82 |
lemmas TFin_is_subset = TFin.dom_subset [THEN subsetD, THEN PowD, standard] |
|
83 |
||
84 |
||
85 |
(** Structural induction on TFin(S,next) **) |
|
86 |
||
87 |
lemma TFin_induct: |
|
88 |
"[| n: TFin(S,next); |
|
89 |
!!x. [| x : TFin(S,next); P(x); next: increasing(S) |] ==> P(next`x); |
|
90 |
!!Y. [| Y <= TFin(S,next); ALL y:Y. P(y) |] ==> P(Union(Y)) |
|
91 |
|] ==> P(n)" |
|
13356 | 92 |
by (erule TFin.induct, blast+) |
13134 | 93 |
|
94 |
||
13356 | 95 |
subsection{*Some Properties of the Transfinite Construction *} |
13134 | 96 |
|
97 |
lemmas increasing_trans = subset_trans [OF _ increasingD2, |
|
98 |
OF _ _ TFin_is_subset] |
|
99 |
||
100 |
(*Lemma 1 of section 3.1*) |
|
101 |
lemma TFin_linear_lemma1: |
|
102 |
"[| n: TFin(S,next); m: TFin(S,next); |
|
103 |
ALL x: TFin(S,next) . x<=m --> x=m | next`x<=m |] |
|
104 |
==> n<=m | next`m<=n" |
|
105 |
apply (erule TFin_induct) |
|
106 |
apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*) |
|
107 |
(*downgrade subsetI from intro! to intro*) |
|
108 |
apply (blast dest: increasing_trans) |
|
109 |
done |
|
110 |
||
111 |
(*Lemma 2 of section 3.2. Interesting in its own right! |
|
112 |
Requires next: increasing(S) in the second induction step. *) |
|
113 |
lemma TFin_linear_lemma2: |
|
114 |
"[| m: TFin(S,next); next: increasing(S) |] |
|
115 |
==> ALL n: TFin(S,next) . n<=m --> n=m | next`n<=m" |
|
116 |
apply (erule TFin_induct) |
|
117 |
apply (rule impI [THEN ballI]) |
|
118 |
(*case split using TFin_linear_lemma1*) |
|
119 |
apply (rule_tac n1 = "n" and m1 = "x" in TFin_linear_lemma1 [THEN disjE], |
|
120 |
assumption+) |
|
121 |
apply (blast del: subsetI |
|
13269 | 122 |
intro: increasing_trans subsetI, blast) |
13134 | 123 |
(*second induction step*) |
124 |
apply (rule impI [THEN ballI]) |
|
125 |
apply (rule Union_lemma0 [THEN disjE]) |
|
126 |
apply (erule_tac [3] disjI2) |
|
13269 | 127 |
prefer 2 apply blast |
13134 | 128 |
apply (rule ballI) |
129 |
apply (drule bspec, assumption) |
|
130 |
apply (drule subsetD, assumption) |
|
131 |
apply (rule_tac n1 = "n" and m1 = "x" in TFin_linear_lemma1 [THEN disjE], |
|
13269 | 132 |
assumption+, blast) |
13134 | 133 |
apply (erule increasingD2 [THEN subset_trans, THEN disjI1]) |
134 |
apply (blast dest: TFin_is_subset)+ |
|
135 |
done |
|
136 |
||
137 |
(*a more convenient form for Lemma 2*) |
|
138 |
lemma TFin_subsetD: |
|
139 |
"[| n<=m; m: TFin(S,next); n: TFin(S,next); next: increasing(S) |] |
|
140 |
==> n=m | next`n<=m" |
|
141 |
by (blast dest: TFin_linear_lemma2 [rule_format]) |
|
142 |
||
143 |
(*Consequences from section 3.3 -- Property 3.2, the ordering is total*) |
|
144 |
lemma TFin_subset_linear: |
|
145 |
"[| m: TFin(S,next); n: TFin(S,next); next: increasing(S) |] |
|
146 |
==> n<=m | m<=n" |
|
147 |
apply (rule disjE) |
|
148 |
apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) |
|
149 |
apply (assumption+, erule disjI2) |
|
150 |
apply (blast del: subsetI |
|
151 |
intro: subsetI increasingD2 [THEN subset_trans] TFin_is_subset) |
|
152 |
done |
|
153 |
||
154 |
||
155 |
(*Lemma 3 of section 3.3*) |
|
156 |
lemma equal_next_upper: |
|
157 |
"[| n: TFin(S,next); m: TFin(S,next); m = next`m |] ==> n<=m" |
|
158 |
apply (erule TFin_induct) |
|
159 |
apply (drule TFin_subsetD) |
|
13269 | 160 |
apply (assumption+, force) |
161 |
apply blast |
|
13134 | 162 |
done |
163 |
||
164 |
(*Property 3.3 of section 3.3*) |
|
165 |
lemma equal_next_Union: "[| m: TFin(S,next); next: increasing(S) |] |
|
166 |
==> m = next`m <-> m = Union(TFin(S,next))" |
|
167 |
apply (rule iffI) |
|
168 |
apply (rule Union_upper [THEN equalityI]) |
|
169 |
apply (rule_tac [2] equal_next_upper [THEN Union_least]) |
|
170 |
apply (assumption+) |
|
171 |
apply (erule ssubst) |
|
13269 | 172 |
apply (rule increasingD2 [THEN equalityI], assumption) |
13134 | 173 |
apply (blast del: subsetI |
174 |
intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+ |
|
175 |
done |
|
176 |
||
177 |
||
13356 | 178 |
subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain*} |
179 |
||
180 |
text{*NOTE: We assume the partial ordering is @{text "\<subseteq>"}, the subset |
|
181 |
relation!*} |
|
13134 | 182 |
|
183 |
(** Defining the "next" operation for Hausdorff's Theorem **) |
|
184 |
||
185 |
lemma chain_subset_Pow: "chain(A) <= Pow(A)" |
|
186 |
apply (unfold chain_def) |
|
187 |
apply (rule Collect_subset) |
|
188 |
done |
|
189 |
||
190 |
lemma super_subset_chain: "super(A,c) <= chain(A)" |
|
191 |
apply (unfold super_def) |
|
192 |
apply (rule Collect_subset) |
|
193 |
done |
|
194 |
||
195 |
lemma maxchain_subset_chain: "maxchain(A) <= chain(A)" |
|
196 |
apply (unfold maxchain_def) |
|
197 |
apply (rule Collect_subset) |
|
198 |
done |
|
199 |
||
200 |
lemma choice_super: "[| ch : (PROD X:Pow(chain(S)) - {0}. X); |
|
201 |
X : chain(S); X ~: maxchain(S) |] |
|
202 |
==> ch ` super(S,X) : super(S,X)" |
|
203 |
apply (erule apply_type) |
|
13269 | 204 |
apply (unfold super_def maxchain_def, blast) |
13134 | 205 |
done |
206 |
||
207 |
lemma choice_not_equals: |
|
208 |
"[| ch : (PROD X:Pow(chain(S)) - {0}. X); |
|
209 |
X : chain(S); X ~: maxchain(S) |] |
|
210 |
==> ch ` super(S,X) ~= X" |
|
211 |
apply (rule notI) |
|
13269 | 212 |
apply (drule choice_super, assumption) |
13134 | 213 |
apply assumption |
214 |
apply (simp add: super_def) |
|
215 |
done |
|
216 |
||
217 |
(*This justifies Definition 4.4*) |
|
218 |
lemma Hausdorff_next_exists: |
|
219 |
"ch: (PROD X: Pow(chain(S))-{0}. X) ==> |
|
220 |
EX next: increasing(S). ALL X: Pow(S). |
|
221 |
next`X = if(X: chain(S)-maxchain(S), ch`super(S,X), X)" |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
222 |
apply (rule_tac x="\<lambda>X\<in>Pow(S). |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
223 |
if X \<in> chain(S) - maxchain(S) then ch ` super(S, X) else X" |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
224 |
in bexI) |
13269 | 225 |
apply force |
13134 | 226 |
apply (unfold increasing_def) |
227 |
apply (rule CollectI) |
|
228 |
apply (rule lam_type) |
|
229 |
apply (simp (no_asm_simp)) |
|
230 |
apply (blast dest: super_subset_chain [THEN subsetD] chain_subset_Pow [THEN subsetD] choice_super) |
|
231 |
(*Now, verify that it increases*) |
|
232 |
apply (simp (no_asm_simp) add: Pow_iff subset_refl) |
|
233 |
apply safe |
|
234 |
apply (drule choice_super) |
|
235 |
apply (assumption+) |
|
13269 | 236 |
apply (simp add: super_def, blast) |
13134 | 237 |
done |
238 |
||
239 |
(*Lemma 4*) |
|
240 |
lemma TFin_chain_lemma4: |
|
241 |
"[| c: TFin(S,next); |
|
242 |
ch: (PROD X: Pow(chain(S))-{0}. X); |
|
243 |
next: increasing(S); |
|
244 |
ALL X: Pow(S). next`X = |
|
245 |
if(X: chain(S)-maxchain(S), ch`super(S,X), X) |] |
|
246 |
==> c: chain(S)" |
|
247 |
apply (erule TFin_induct) |
|
248 |
apply (simp (no_asm_simp) add: chain_subset_Pow [THEN subsetD, THEN PowD] |
|
249 |
choice_super [THEN super_subset_chain [THEN subsetD]]) |
|
250 |
apply (unfold chain_def) |
|
13269 | 251 |
apply (rule CollectI, blast, safe) |
13356 | 252 |
apply (rule_tac m1=B and n1=Ba in TFin_subset_linear [THEN disjE], fast+) |
253 |
(*Blast_tac's slow*) |
|
13134 | 254 |
done |
255 |
||
13356 | 256 |
theorem Hausdorff: "EX c. c : maxchain(S)" |
13134 | 257 |
apply (rule AC_Pi_Pow [THEN exE]) |
13269 | 258 |
apply (rule Hausdorff_next_exists [THEN bexE], assumption) |
13134 | 259 |
apply (rename_tac ch "next") |
260 |
apply (subgoal_tac "Union (TFin (S,next)) : chain (S) ") |
|
261 |
prefer 2 |
|
262 |
apply (blast intro!: TFin_chain_lemma4 subset_refl [THEN TFin_UnionI]) |
|
263 |
apply (rule_tac x = "Union (TFin (S,next))" in exI) |
|
264 |
apply (rule classical) |
|
265 |
apply (subgoal_tac "next ` Union (TFin (S,next)) = Union (TFin (S,next))") |
|
266 |
apply (rule_tac [2] equal_next_Union [THEN iffD2, symmetric]) |
|
267 |
apply (rule_tac [2] subset_refl [THEN TFin_UnionI]) |
|
13269 | 268 |
prefer 2 apply assumption |
13134 | 269 |
apply (rule_tac [2] refl) |
270 |
apply (simp add: subset_refl [THEN TFin_UnionI, |
|
271 |
THEN TFin.dom_subset [THEN subsetD, THEN PowD]]) |
|
272 |
apply (erule choice_not_equals [THEN notE]) |
|
273 |
apply (assumption+) |
|
274 |
done |
|
275 |
||
276 |
||
13356 | 277 |
subsection{*Zorn's Lemma*} |
278 |
||
279 |
text{*If all chains in S have upper bounds in S, |
|
280 |
then S contains a maximal element *} |
|
13134 | 281 |
|
282 |
(*Used in the proof of Zorn's Lemma*) |
|
283 |
lemma chain_extend: |
|
284 |
"[| c: chain(A); z: A; ALL x:c. x<=z |] ==> cons(z,c) : chain(A)" |
|
13356 | 285 |
by (unfold chain_def, blast) |
13134 | 286 |
|
287 |
lemma Zorn: "ALL c: chain(S). Union(c) : S ==> EX y:S. ALL z:S. y<=z --> y=z" |
|
288 |
apply (rule Hausdorff [THEN exE]) |
|
289 |
apply (simp add: maxchain_def) |
|
290 |
apply (rename_tac c) |
|
291 |
apply (rule_tac x = "Union (c)" in bexI) |
|
13269 | 292 |
prefer 2 apply blast |
13134 | 293 |
apply safe |
294 |
apply (rename_tac z) |
|
295 |
apply (rule classical) |
|
296 |
apply (subgoal_tac "cons (z,c) : super (S,c) ") |
|
297 |
apply (blast elim: equalityE) |
|
13269 | 298 |
apply (unfold super_def, safe) |
13134 | 299 |
apply (fast elim: chain_extend) |
300 |
apply (fast elim: equalityE) |
|
301 |
done |
|
302 |
||
303 |
||
13356 | 304 |
subsection{*Zermelo's Theorem: Every Set can be Well-Ordered*} |
13134 | 305 |
|
306 |
(*Lemma 5*) |
|
307 |
lemma TFin_well_lemma5: |
|
308 |
"[| n: TFin(S,next); Z <= TFin(S,next); z:Z; ~ Inter(Z) : Z |] |
|
309 |
==> ALL m:Z. n<=m" |
|
310 |
apply (erule TFin_induct) |
|
13269 | 311 |
prefer 2 apply blast (*second induction step is easy*) |
13134 | 312 |
apply (rule ballI) |
13269 | 313 |
apply (rule bspec [THEN TFin_subsetD, THEN disjE], auto) |
13134 | 314 |
apply (subgoal_tac "m = Inter (Z) ") |
315 |
apply blast+ |
|
316 |
done |
|
317 |
||
318 |
(*Well-ordering of TFin(S,next)*) |
|
319 |
lemma well_ord_TFin_lemma: "[| Z <= TFin(S,next); z:Z |] ==> Inter(Z) : Z" |
|
320 |
apply (rule classical) |
|
321 |
apply (subgoal_tac "Z = {Union (TFin (S,next))}") |
|
322 |
apply (simp (no_asm_simp) add: Inter_singleton) |
|
323 |
apply (erule equal_singleton) |
|
324 |
apply (rule Union_upper [THEN equalityI]) |
|
13269 | 325 |
apply (rule_tac [2] subset_refl [THEN TFin_UnionI, THEN TFin_well_lemma5, THEN bspec], blast+) |
13134 | 326 |
done |
327 |
||
328 |
(*This theorem just packages the previous result*) |
|
329 |
lemma well_ord_TFin: |
|
330 |
"next: increasing(S) ==> well_ord(TFin(S,next), Subset_rel(TFin(S,next)))" |
|
331 |
apply (rule well_ordI) |
|
332 |
apply (unfold Subset_rel_def linear_def) |
|
333 |
(*Prove the well-foundedness goal*) |
|
334 |
apply (rule wf_onI) |
|
13269 | 335 |
apply (frule well_ord_TFin_lemma, assumption) |
336 |
apply (drule_tac x = "Inter (Z) " in bspec, assumption) |
|
13134 | 337 |
apply blast |
338 |
(*Now prove the linearity goal*) |
|
339 |
apply (intro ballI) |
|
340 |
apply (case_tac "x=y") |
|
13269 | 341 |
apply blast |
13134 | 342 |
(*The x~=y case remains*) |
343 |
apply (rule_tac n1=x and m1=y in TFin_subset_linear [THEN disjE], |
|
13269 | 344 |
assumption+, blast+) |
13134 | 345 |
done |
346 |
||
347 |
(** Defining the "next" operation for Zermelo's Theorem **) |
|
348 |
||
349 |
lemma choice_Diff: |
|
350 |
"[| ch \<in> (\<Pi>X \<in> Pow(S) - {0}. X); X \<subseteq> S; X\<noteq>S |] ==> ch ` (S-X) \<in> S-X" |
|
351 |
apply (erule apply_type) |
|
352 |
apply (blast elim!: equalityE) |
|
353 |
done |
|
354 |
||
355 |
(*This justifies Definition 6.1*) |
|
356 |
lemma Zermelo_next_exists: |
|
357 |
"ch: (PROD X: Pow(S)-{0}. X) ==> |
|
358 |
EX next: increasing(S). ALL X: Pow(S). |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
359 |
next`X = (if X=S then S else cons(ch`(S-X), X))" |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
360 |
apply (rule_tac x="\<lambda>X\<in>Pow(S). if X=S then S else cons(ch`(S-X), X)" |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
13134
diff
changeset
|
361 |
in bexI) |
13269 | 362 |
apply force |
13134 | 363 |
apply (unfold increasing_def) |
364 |
apply (rule CollectI) |
|
365 |
apply (rule lam_type) |
|
366 |
(*Type checking is surprisingly hard!*) |
|
367 |
apply (simp (no_asm_simp) add: Pow_iff cons_subset_iff subset_refl) |
|
368 |
apply (blast intro!: choice_Diff [THEN DiffD1]) |
|
369 |
(*Verify that it increases*) |
|
370 |
apply (intro allI impI) |
|
371 |
apply (simp add: Pow_iff subset_consI subset_refl) |
|
372 |
done |
|
373 |
||
374 |
||
375 |
(*The construction of the injection*) |
|
376 |
lemma choice_imp_injection: |
|
377 |
"[| 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 |
apply (rule_tac d = "%y. ch` (S-y) " in lam_injective) |
|
383 |
apply (rule DiffI) |
|
384 |
apply (rule Collect_subset [THEN TFin_UnionI]) |
|
385 |
apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE) |
|
386 |
apply (subgoal_tac "x ~: Union ({y: TFin (S,next) . x~: y}) ") |
|
387 |
prefer 2 apply (blast elim: equalityE) |
|
388 |
apply (subgoal_tac "Union ({y: TFin (S,next) . x~: y}) ~= S") |
|
389 |
prefer 2 apply (blast elim: equalityE) |
|
390 |
(*For proving x : next`Union(...) |
|
391 |
Abrial & Laffitte's justification appears to be faulty.*) |
|
392 |
apply (subgoal_tac "~ next ` Union ({y: TFin (S,next) . x~: y}) <= Union ({y: TFin (S,next) . x~: y}) ") |
|
393 |
prefer 2 |
|
394 |
apply (simp del: Union_iff |
|
395 |
add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset] |
|
396 |
Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2]) |
|
397 |
apply (subgoal_tac "x : next ` Union ({y: TFin (S,next) . x~: y}) ") |
|
398 |
prefer 2 |
|
399 |
apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI) |
|
400 |
(*End of the lemmas!*) |
|
401 |
apply (simp add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]) |
|
402 |
done |
|
403 |
||
404 |
(*The wellordering theorem*) |
|
13356 | 405 |
theorem AC_well_ord: "EX r. well_ord(S,r)" |
13134 | 406 |
apply (rule AC_Pi_Pow [THEN exE]) |
13269 | 407 |
apply (rule Zermelo_next_exists [THEN bexE], assumption) |
13134 | 408 |
apply (rule exI) |
409 |
apply (rule well_ord_rvimage) |
|
410 |
apply (erule_tac [2] well_ord_TFin) |
|
13269 | 411 |
apply (rule choice_imp_injection [THEN inj_weaken_type], blast+) |
13134 | 412 |
done |
516 | 413 |
|
414 |
end |