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