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