| author | wenzelm | 
| Mon, 18 Dec 2023 21:46:51 +0100 | |
| changeset 79283 | 2c5d4b4ea3a2 | 
| parent 76217 | 8655344f1cf6 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/Zorn.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 516 | 3 | Copyright 1994 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section\<open>Zorn's Lemma\<close> | 
| 13356 | 7 | |
| 68490 
eb53f944c8cd
simplified ZF theory names (in contrast to 6a0801279f4c): session-qualification already achieves disjointness;
 wenzelm parents: 
67443diff
changeset | 8 | theory Zorn imports OrderArith AC Inductive begin | 
| 516 | 9 | |
| 60770 | 10 | text\<open>Based upon the unpublished article ``Towards the Mechanization of the | 
| 11 | Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.\<close> | |
| 13356 | 12 | |
| 24893 | 13 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 14 | Subset_rel :: "i\<Rightarrow>i" where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 15 |    "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=\<langle>x,y\<rangle> \<and> x<=y \<and> x\<noteq>y}"
 | 
| 13134 | 16 | |
| 24893 | 17 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 18 | chain :: "i\<Rightarrow>i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 19 |    "chain(A)      \<equiv> {F \<in> Pow(A). \<forall>X\<in>F. \<forall>Y\<in>F. X<=Y | Y<=X}"
 | 
| 516 | 20 | |
| 24893 | 21 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 22 | super :: "[i,i]\<Rightarrow>i" where | 
| 76214 | 23 |    "super(A,c)    \<equiv> {d \<in> chain(A). c<=d \<and> c\<noteq>d}"
 | 
| 14653 | 24 | |
| 24893 | 25 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 26 | maxchain :: "i\<Rightarrow>i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 27 |    "maxchain(A)   \<equiv> {c \<in> chain(A). super(A,c)=0}"
 | 
| 13558 | 28 | |
| 24893 | 29 | definition | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 30 | increasing :: "i\<Rightarrow>i" where | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 31 |     "increasing(A) \<equiv> {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A \<longrightarrow> x<=f`x}"
 | 
| 516 | 32 | |
| 13356 | 33 | |
| 60770 | 34 | text\<open>Lemma for the inductive definition below\<close> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 35 | lemma Union_in_Pow: "Y \<in> Pow(Pow(A)) \<Longrightarrow> \<Union>(Y) \<in> Pow(A)" | 
| 13356 | 36 | by blast | 
| 37 | ||
| 38 | ||
| 60770 | 39 | text\<open>We could make the inductive definition conditional on | 
| 69593 | 40 | \<^term>\<open>next \<in> increasing(S)\<close> | 
| 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 | |
| 60770 | 43 | are therefore unconditional.\<close> | 
| 516 | 44 | consts | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 45 | "TFin" :: "[i,i]\<Rightarrow>i" | 
| 516 | 46 | |
| 47 | inductive | |
| 46820 | 48 | domains "TFin(S,next)" \<subseteq> "Pow(S)" | 
| 13134 | 49 | intros | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 50 | nextI: "\<lbrakk>x \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 51 | \<Longrightarrow> next`x \<in> TFin(S,next)" | 
| 516 | 52 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 53 | Pow_UnionI: "Y \<in> Pow(TFin(S,next)) \<Longrightarrow> \<Union>(Y) \<in> TFin(S,next)" | 
| 516 | 54 | |
| 6053 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1478diff
changeset | 55 | monos Pow_mono | 
| 
8a1059aa01f0
new inductive, datatype and primrec packages, etc.
 paulson parents: 
1478diff
changeset | 56 | con_defs increasing_def | 
| 13134 | 57 | type_intros CollectD1 [THEN apply_funtype] Union_in_Pow | 
| 58 | ||
| 59 | ||
| 60770 | 60 | subsection\<open>Mathematical Preamble\<close> | 
| 13134 | 61 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 62 | lemma Union_lemma0: "(\<forall>x\<in>C. x<=A | B<=x) \<Longrightarrow> \<Union>(C)<=A | B<=\<Union>(C)" | 
| 13269 | 63 | by blast | 
| 13134 | 64 | |
| 13356 | 65 | lemma Inter_lemma0: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 66 | "\<lbrakk>c \<in> C; \<forall>x\<in>C. A<=x | x<=B\<rbrakk> \<Longrightarrow> A \<subseteq> \<Inter>(C) | \<Inter>(C) \<subseteq> B" | 
| 13269 | 67 | by blast | 
| 13134 | 68 | |
| 69 | ||
| 60770 | 70 | subsection\<open>The Transfinite Construction\<close> | 
| 13134 | 71 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 72 | lemma increasingD1: "f \<in> increasing(A) \<Longrightarrow> f \<in> Pow(A)->Pow(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 73 | unfolding increasing_def | 
| 13134 | 74 | apply (erule CollectD1) | 
| 75 | done | |
| 76 | ||
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 77 | lemma increasingD2: "\<lbrakk>f \<in> increasing(A); x<=A\<rbrakk> \<Longrightarrow> x \<subseteq> f`x" | 
| 13269 | 78 | by (unfold increasing_def, blast) | 
| 13134 | 79 | |
| 45602 | 80 | lemmas TFin_UnionI = PowI [THEN TFin.Pow_UnionI] | 
| 13134 | 81 | |
| 45602 | 82 | lemmas TFin_is_subset = TFin.dom_subset [THEN subsetD, THEN PowD] | 
| 13134 | 83 | |
| 84 | ||
| 69593 | 85 | text\<open>Structural induction on \<^term>\<open>TFin(S,next)\<close>\<close> | 
| 13134 | 86 | lemma TFin_induct: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 87 | "\<lbrakk>n \<in> TFin(S,next); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 88 | \<And>x. \<lbrakk>x \<in> TFin(S,next); P(x); next \<in> increasing(S)\<rbrakk> \<Longrightarrow> P(next`x); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 89 | \<And>Y. \<lbrakk>Y \<subseteq> TFin(S,next); \<forall>y\<in>Y. P(y)\<rbrakk> \<Longrightarrow> P(\<Union>(Y)) | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 90 | \<rbrakk> \<Longrightarrow> P(n)" | 
| 13356 | 91 | by (erule TFin.induct, blast+) | 
| 13134 | 92 | |
| 93 | ||
| 60770 | 94 | subsection\<open>Some Properties of the Transfinite Construction\<close> | 
| 13134 | 95 | |
| 13558 | 96 | lemmas increasing_trans = subset_trans [OF _ increasingD2, | 
| 13134 | 97 | OF _ _ TFin_is_subset] | 
| 98 | ||
| 60770 | 99 | text\<open>Lemma 1 of section 3.1\<close> | 
| 13134 | 100 | lemma TFin_linear_lemma1: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 101 | "\<lbrakk>n \<in> TFin(S,next); m \<in> TFin(S,next); | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 102 | \<forall>x \<in> TFin(S,next) . x<=m \<longrightarrow> x=m | next`x<=m\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 103 | \<Longrightarrow> n<=m | next`m<=n" | 
| 13134 | 104 | apply (erule TFin_induct) | 
| 105 | apply (erule_tac [2] Union_lemma0) (*or just Blast_tac*) | |
| 106 | (*downgrade subsetI from intro! to intro*) | |
| 107 | apply (blast dest: increasing_trans) | |
| 108 | done | |
| 109 | ||
| 60770 | 110 | text\<open>Lemma 2 of section 3.2. Interesting in its own right! | 
| 69593 | 111 | Requires \<^term>\<open>next \<in> increasing(S)\<close> in the second induction step.\<close> | 
| 13134 | 112 | lemma TFin_linear_lemma2: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 113 | "\<lbrakk>m \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 114 | \<Longrightarrow> \<forall>n \<in> TFin(S,next). n<=m \<longrightarrow> n=m | next`n \<subseteq> m" | 
| 13134 | 115 | apply (erule TFin_induct) | 
| 116 | apply (rule impI [THEN ballI]) | |
| 61798 | 117 | txt\<open>case split using \<open>TFin_linear_lemma1\<close>\<close> | 
| 13784 | 118 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | 
| 13134 | 119 | assumption+) | 
| 120 | apply (blast del: subsetI | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 121 | intro: increasing_trans subsetI, blast) | 
| 60770 | 122 | txt\<open>second induction step\<close> | 
| 13134 | 123 | apply (rule impI [THEN ballI]) | 
| 124 | apply (rule Union_lemma0 [THEN disjE]) | |
| 125 | apply (erule_tac [3] disjI2) | |
| 13558 | 126 | prefer 2 apply blast | 
| 13134 | 127 | apply (rule ballI) | 
| 13558 | 128 | apply (drule bspec, assumption) | 
| 129 | apply (drule subsetD, assumption) | |
| 13784 | 130 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | 
| 13558 | 131 | assumption+, blast) | 
| 13134 | 132 | apply (erule increasingD2 [THEN subset_trans, THEN disjI1]) | 
| 133 | apply (blast dest: TFin_is_subset)+ | |
| 134 | done | |
| 135 | ||
| 60770 | 136 | text\<open>a more convenient form for Lemma 2\<close> | 
| 13134 | 137 | lemma TFin_subsetD: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 138 | "\<lbrakk>n<=m; m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 139 | \<Longrightarrow> n=m | next`n \<subseteq> m" | 
| 13558 | 140 | by (blast dest: TFin_linear_lemma2 [rule_format]) | 
| 13134 | 141 | |
| 60770 | 142 | text\<open>Consequences from section 3.3 -- Property 3.2, the ordering is total\<close> | 
| 13134 | 143 | lemma TFin_subset_linear: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 144 | "\<lbrakk>m \<in> TFin(S,next); n \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 145 | \<Longrightarrow> n \<subseteq> m | m<=n" | 
| 13558 | 146 | apply (rule disjE) | 
| 13134 | 147 | apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) | 
| 148 | apply (assumption+, erule disjI2) | |
| 13558 | 149 | apply (blast del: subsetI | 
| 13134 | 150 | intro: subsetI increasingD2 [THEN subset_trans] TFin_is_subset) | 
| 151 | done | |
| 152 | ||
| 153 | ||
| 60770 | 154 | text\<open>Lemma 3 of section 3.3\<close> | 
| 13134 | 155 | lemma equal_next_upper: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 156 | "\<lbrakk>n \<in> TFin(S,next); m \<in> TFin(S,next); m = next`m\<rbrakk> \<Longrightarrow> n \<subseteq> m" | 
| 13134 | 157 | apply (erule TFin_induct) | 
| 158 | apply (drule TFin_subsetD) | |
| 13784 | 159 | apply (assumption+, force, blast) | 
| 13134 | 160 | done | 
| 161 | ||
| 60770 | 162 | text\<open>Property 3.3 of section 3.3\<close> | 
| 13558 | 163 | lemma equal_next_Union: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 164 | "\<lbrakk>m \<in> TFin(S,next); next \<in> increasing(S)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 165 | \<Longrightarrow> m = next`m <-> m = \<Union>(TFin(S,next))" | 
| 13134 | 166 | apply (rule iffI) | 
| 167 | apply (rule Union_upper [THEN equalityI]) | |
| 168 | apply (rule_tac [2] equal_next_upper [THEN Union_least]) | |
| 169 | apply (assumption+) | |
| 170 | apply (erule ssubst) | |
| 13269 | 171 | apply (rule increasingD2 [THEN equalityI], assumption) | 
| 13134 | 172 | apply (blast del: subsetI | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 173 | intro: subsetI TFin_UnionI TFin.nextI TFin_is_subset)+ | 
| 13134 | 174 | done | 
| 175 | ||
| 176 | ||
| 60770 | 177 | subsection\<open>Hausdorff's Theorem: Every Set Contains a Maximal Chain\<close> | 
| 13356 | 178 | |
| 61798 | 179 | text\<open>NOTE: We assume the partial ordering is \<open>\<subseteq>\<close>, the subset | 
| 60770 | 180 | relation!\<close> | 
| 13134 | 181 | |
| 60770 | 182 | text\<open>* Defining the "next" operation for Hausdorff's Theorem *\<close> | 
| 13134 | 183 | |
| 46820 | 184 | lemma chain_subset_Pow: "chain(A) \<subseteq> Pow(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 185 | unfolding chain_def | 
| 13134 | 186 | apply (rule Collect_subset) | 
| 187 | done | |
| 188 | ||
| 46820 | 189 | lemma super_subset_chain: "super(A,c) \<subseteq> chain(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 190 | unfolding super_def | 
| 13134 | 191 | apply (rule Collect_subset) | 
| 192 | done | |
| 193 | ||
| 46820 | 194 | lemma maxchain_subset_chain: "maxchain(A) \<subseteq> chain(A)" | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 195 | unfolding maxchain_def | 
| 13134 | 196 | apply (rule Collect_subset) | 
| 197 | done | |
| 198 | ||
| 13558 | 199 | lemma choice_super: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 200 |      "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S);  X \<notin> maxchain(S)\<rbrakk>
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 201 | \<Longrightarrow> ch ` super(S,X) \<in> super(S,X)" | 
| 13134 | 202 | apply (erule apply_type) | 
| 13269 | 203 | apply (unfold super_def maxchain_def, blast) | 
| 13134 | 204 | done | 
| 205 | ||
| 206 | lemma choice_not_equals: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 207 |      "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(chain(S)) - {0}. X); X \<in> chain(S);  X \<notin> maxchain(S)\<rbrakk>
 | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 208 | \<Longrightarrow> ch ` super(S,X) \<noteq> X" | 
| 13134 | 209 | apply (rule notI) | 
| 13784 | 210 | apply (drule choice_super, assumption, assumption) | 
| 13134 | 211 | apply (simp add: super_def) | 
| 212 | done | |
| 213 | ||
| 60770 | 214 | text\<open>This justifies Definition 4.4\<close> | 
| 13134 | 215 | lemma Hausdorff_next_exists: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 216 |      "ch \<in> (\<Prod>X \<in> Pow(chain(S))-{0}. X) \<Longrightarrow>
 | 
| 13558 | 217 | \<exists>next \<in> increasing(S). \<forall>X \<in> Pow(S). | 
| 218 | next`X = if(X \<in> chain(S)-maxchain(S), ch`super(S,X), X)" | |
| 219 | apply (rule_tac x="\<lambda>X\<in>Pow(S). | |
| 220 | 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 | 221 | in bexI) | 
| 13558 | 222 | apply force | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 223 | unfolding increasing_def | 
| 13134 | 224 | apply (rule CollectI) | 
| 225 | apply (rule lam_type) | |
| 226 | apply (simp (no_asm_simp)) | |
| 13558 | 227 | apply (blast dest: super_subset_chain [THEN subsetD] | 
| 228 | chain_subset_Pow [THEN subsetD] choice_super) | |
| 60770 | 229 | txt\<open>Now, verify that it increases\<close> | 
| 13134 | 230 | apply (simp (no_asm_simp) add: Pow_iff subset_refl) | 
| 231 | apply safe | |
| 232 | apply (drule choice_super) | |
| 233 | apply (assumption+) | |
| 13269 | 234 | apply (simp add: super_def, blast) | 
| 13134 | 235 | done | 
| 236 | ||
| 60770 | 237 | text\<open>Lemma 4\<close> | 
| 13134 | 238 | lemma TFin_chain_lemma4: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 239 | "\<lbrakk>c \<in> TFin(S,next); | 
| 61980 | 240 |          ch \<in> (\<Prod>X \<in> Pow(chain(S))-{0}. X);
 | 
| 13558 | 241 | next \<in> increasing(S); | 
| 242 | \<forall>X \<in> Pow(S). next`X = | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 243 | if(X \<in> chain(S)-maxchain(S), ch`super(S,X), X)\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 244 | \<Longrightarrow> c \<in> chain(S)" | 
| 13134 | 245 | apply (erule TFin_induct) | 
| 13558 | 246 | apply (simp (no_asm_simp) add: chain_subset_Pow [THEN subsetD, THEN PowD] | 
| 13134 | 247 | choice_super [THEN super_subset_chain [THEN subsetD]]) | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 248 | unfolding chain_def | 
| 13269 | 249 | apply (rule CollectI, blast, safe) | 
| 13558 | 250 | apply (rule_tac m1=B and n1=Ba in TFin_subset_linear [THEN disjE], fast+) | 
| 61798 | 251 | txt\<open>\<open>Blast_tac's\<close> slow\<close> | 
| 13134 | 252 | done | 
| 253 | ||
| 13558 | 254 | theorem Hausdorff: "\<exists>c. c \<in> maxchain(S)" | 
| 13134 | 255 | apply (rule AC_Pi_Pow [THEN exE]) | 
| 13269 | 256 | apply (rule Hausdorff_next_exists [THEN bexE], assumption) | 
| 13134 | 257 | apply (rename_tac ch "next") | 
| 46820 | 258 | apply (subgoal_tac "\<Union>(TFin (S,next)) \<in> chain (S) ") | 
| 13134 | 259 | prefer 2 | 
| 260 | apply (blast intro!: TFin_chain_lemma4 subset_refl [THEN TFin_UnionI]) | |
| 46820 | 261 | apply (rule_tac x = "\<Union>(TFin (S,next))" in exI) | 
| 13134 | 262 | apply (rule classical) | 
| 46820 | 263 | apply (subgoal_tac "next ` Union(TFin (S,next)) = \<Union>(TFin (S,next))") | 
| 13134 | 264 | apply (rule_tac [2] equal_next_Union [THEN iffD2, symmetric]) | 
| 265 | apply (rule_tac [2] subset_refl [THEN TFin_UnionI]) | |
| 13269 | 266 | prefer 2 apply assumption | 
| 13134 | 267 | apply (rule_tac [2] refl) | 
| 13558 | 268 | apply (simp add: subset_refl [THEN TFin_UnionI, | 
| 13134 | 269 | THEN TFin.dom_subset [THEN subsetD, THEN PowD]]) | 
| 270 | apply (erule choice_not_equals [THEN notE]) | |
| 271 | apply (assumption+) | |
| 272 | done | |
| 273 | ||
| 274 | ||
| 60770 | 275 | subsection\<open>Zorn's Lemma: If All Chains in S Have Upper Bounds In S, | 
| 276 | then S contains a Maximal Element\<close> | |
| 13356 | 277 | |
| 60770 | 278 | text\<open>Used in the proof of Zorn's Lemma\<close> | 
| 13558 | 279 | lemma chain_extend: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 280 | "\<lbrakk>c \<in> chain(A); z \<in> A; \<forall>x \<in> c. x<=z\<rbrakk> \<Longrightarrow> cons(z,c) \<in> chain(A)" | 
| 13356 | 281 | by (unfold chain_def, blast) | 
| 13134 | 282 | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 283 | lemma Zorn: "\<forall>c \<in> chain(S). \<Union>(c) \<in> S \<Longrightarrow> \<exists>y \<in> S. \<forall>z \<in> S. y<=z \<longrightarrow> y=z" | 
| 13134 | 284 | apply (rule Hausdorff [THEN exE]) | 
| 285 | apply (simp add: maxchain_def) | |
| 286 | apply (rename_tac c) | |
| 46820 | 287 | apply (rule_tac x = "\<Union>(c)" in bexI) | 
| 13269 | 288 | prefer 2 apply blast | 
| 13134 | 289 | apply safe | 
| 290 | apply (rename_tac z) | |
| 291 | apply (rule classical) | |
| 13558 | 292 | apply (subgoal_tac "cons (z,c) \<in> super (S,c) ") | 
| 13134 | 293 | apply (blast elim: equalityE) | 
| 13269 | 294 | apply (unfold super_def, safe) | 
| 13134 | 295 | apply (fast elim: chain_extend) | 
| 296 | apply (fast elim: equalityE) | |
| 297 | done | |
| 298 | ||
| 60770 | 299 | text \<open>Alternative version of Zorn's Lemma\<close> | 
| 27704 | 300 | |
| 301 | theorem Zorn2: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 302 | "\<forall>c \<in> chain(S). \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y \<Longrightarrow> \<exists>y \<in> S. \<forall>z \<in> S. y<=z \<longrightarrow> y=z" | 
| 27704 | 303 | apply (cut_tac Hausdorff maxchain_subset_chain) | 
| 304 | apply (erule exE) | |
| 305 | apply (drule subsetD, assumption) | |
| 306 | apply (drule bspec, assumption, erule bexE) | |
| 307 | apply (rule_tac x = y in bexI) | |
| 308 | prefer 2 apply assumption | |
| 309 | apply clarify | |
| 310 | apply rule apply assumption | |
| 311 | apply rule | |
| 312 | apply (rule ccontr) | |
| 313 | apply (frule_tac z=z in chain_extend) | |
| 314 | apply (assumption, blast) | |
| 76217 | 315 | unfolding maxchain_def super_def | 
| 27704 | 316 | apply (blast elim!: equalityCE) | 
| 317 | done | |
| 318 | ||
| 13134 | 319 | |
| 60770 | 320 | subsection\<open>Zermelo's Theorem: Every Set can be Well-Ordered\<close> | 
| 13134 | 321 | |
| 60770 | 322 | text\<open>Lemma 5\<close> | 
| 13134 | 323 | lemma TFin_well_lemma5: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 324 | "\<lbrakk>n \<in> TFin(S,next); Z \<subseteq> TFin(S,next); z:Z; \<not> \<Inter>(Z) \<in> Z\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 325 | \<Longrightarrow> \<forall>m \<in> Z. n \<subseteq> m" | 
| 13134 | 326 | apply (erule TFin_induct) | 
| 60770 | 327 | prefer 2 apply blast txt\<open>second induction step is easy\<close> | 
| 13134 | 328 | apply (rule ballI) | 
| 13558 | 329 | apply (rule bspec [THEN TFin_subsetD, THEN disjE], auto) | 
| 46820 | 330 | apply (subgoal_tac "m = \<Inter>(Z) ") | 
| 13134 | 331 | apply blast+ | 
| 332 | done | |
| 333 | ||
| 69593 | 334 | text\<open>Well-ordering of \<^term>\<open>TFin(S,next)\<close>\<close> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 335 | lemma well_ord_TFin_lemma: "\<lbrakk>Z \<subseteq> TFin(S,next); z \<in> Z\<rbrakk> \<Longrightarrow> \<Inter>(Z) \<in> Z" | 
| 13134 | 336 | apply (rule classical) | 
| 46820 | 337 | apply (subgoal_tac "Z = {\<Union>(TFin (S,next))}")
 | 
| 13134 | 338 | apply (simp (no_asm_simp) add: Inter_singleton) | 
| 339 | apply (erule equal_singleton) | |
| 340 | apply (rule Union_upper [THEN equalityI]) | |
| 13269 | 341 | apply (rule_tac [2] subset_refl [THEN TFin_UnionI, THEN TFin_well_lemma5, THEN bspec], blast+) | 
| 13134 | 342 | done | 
| 343 | ||
| 60770 | 344 | text\<open>This theorem just packages the previous result\<close> | 
| 13134 | 345 | lemma well_ord_TFin: | 
| 13558 | 346 | "next \<in> increasing(S) | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 347 | \<Longrightarrow> well_ord(TFin(S,next), Subset_rel(TFin(S,next)))" | 
| 13134 | 348 | apply (rule well_ordI) | 
| 76217 | 349 | unfolding Subset_rel_def linear_def | 
| 60770 | 350 | txt\<open>Prove the well-foundedness goal\<close> | 
| 13134 | 351 | apply (rule wf_onI) | 
| 13269 | 352 | apply (frule well_ord_TFin_lemma, assumption) | 
| 46820 | 353 | apply (drule_tac x = "\<Inter>(Z) " in bspec, assumption) | 
| 13134 | 354 | apply blast | 
| 60770 | 355 | txt\<open>Now prove the linearity goal\<close> | 
| 13134 | 356 | apply (intro ballI) | 
| 357 | apply (case_tac "x=y") | |
| 13269 | 358 | apply blast | 
| 69593 | 359 | txt\<open>The \<^term>\<open>x\<noteq>y\<close> case remains\<close> | 
| 13134 | 360 | apply (rule_tac n1=x and m1=y in TFin_subset_linear [THEN disjE], | 
| 13269 | 361 | assumption+, blast+) | 
| 13134 | 362 | done | 
| 363 | ||
| 60770 | 364 | text\<open>* Defining the "next" operation for Zermelo's Theorem *\<close> | 
| 13134 | 365 | |
| 366 | lemma choice_Diff: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 367 |      "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(S) - {0}. X);  X \<subseteq> S;  X\<noteq>S\<rbrakk> \<Longrightarrow> ch ` (S-X) \<in> S-X"
 | 
| 13134 | 368 | apply (erule apply_type) | 
| 369 | apply (blast elim!: equalityE) | |
| 370 | done | |
| 371 | ||
| 60770 | 372 | text\<open>This justifies Definition 6.1\<close> | 
| 13134 | 373 | lemma Zermelo_next_exists: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 374 |      "ch \<in> (\<Prod>X \<in> Pow(S)-{0}. X) \<Longrightarrow>
 | 
| 13558 | 375 | \<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 | 376 | 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 | 377 | 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 | 378 | in bexI) | 
| 13558 | 379 | apply force | 
| 76216 
9fc34f76b4e8
getting rid of apply (unfold ...)
 paulson <lp15@cam.ac.uk> parents: 
76215diff
changeset | 380 | unfolding increasing_def | 
| 13134 | 381 | apply (rule CollectI) | 
| 382 | apply (rule lam_type) | |
| 60770 | 383 | txt\<open>Type checking is surprisingly hard!\<close> | 
| 13134 | 384 | apply (simp (no_asm_simp) add: Pow_iff cons_subset_iff subset_refl) | 
| 385 | apply (blast intro!: choice_Diff [THEN DiffD1]) | |
| 60770 | 386 | txt\<open>Verify that it increases\<close> | 
| 13558 | 387 | apply (intro allI impI) | 
| 13134 | 388 | apply (simp add: Pow_iff subset_consI subset_refl) | 
| 389 | done | |
| 390 | ||
| 391 | ||
| 60770 | 392 | text\<open>The construction of the injection\<close> | 
| 13134 | 393 | lemma choice_imp_injection: | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 394 |      "\<lbrakk>ch \<in> (\<Prod>X \<in> Pow(S)-{0}. X);
 | 
| 13558 | 395 | next \<in> increasing(S); | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 396 | \<forall>X \<in> Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X))\<rbrakk> | 
| 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 397 |       \<Longrightarrow> (\<lambda> x \<in> S. \<Union>({y \<in> TFin(S,next). x \<notin> y}))
 | 
| 13558 | 398 |                \<in> inj(S, TFin(S,next) - {S})"
 | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 399 | apply (rule_tac d = "\<lambda>y. ch` (S-y) " in lam_injective) | 
| 13134 | 400 | apply (rule DiffI) | 
| 401 | apply (rule Collect_subset [THEN TFin_UnionI]) | |
| 402 | apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE) | |
| 46820 | 403 | apply (subgoal_tac "x \<notin> \<Union>({y \<in> TFin (S,next) . x \<notin> y}) ")
 | 
| 13134 | 404 | prefer 2 apply (blast elim: equalityE) | 
| 46820 | 405 | apply (subgoal_tac "\<Union>({y \<in> TFin (S,next) . x \<notin> y}) \<noteq> S")
 | 
| 13134 | 406 | prefer 2 apply (blast elim: equalityE) | 
| 61798 | 407 | txt\<open>For proving \<open>x \<in> next`\<Union>(...)\<close>. | 
| 60770 | 408 | Abrial and Laffitte's justification appears to be faulty.\<close> | 
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 409 | apply (subgoal_tac "\<not> next ` Union({y \<in> TFin (S,next) . x \<notin> y}) 
 | 
| 46820 | 410 |                     \<subseteq> \<Union>({y \<in> TFin (S,next) . x \<notin> y}) ")
 | 
| 13558 | 411 | prefer 2 | 
| 412 | apply (simp del: Union_iff | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 413 | add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset] | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 414 | Pow_iff cons_subset_iff subset_refl choice_Diff [THEN DiffD2]) | 
| 46820 | 415 | apply (subgoal_tac "x \<in> next ` Union({y \<in> TFin (S,next) . x \<notin> y}) ")
 | 
| 13558 | 416 | prefer 2 | 
| 417 | apply (blast intro!: Collect_subset [THEN TFin_UnionI] TFin.nextI) | |
| 60770 | 418 | txt\<open>End of the lemmas!\<close> | 
| 13134 | 419 | apply (simp add: Collect_subset [THEN TFin_UnionI, THEN TFin_is_subset]) | 
| 420 | done | |
| 421 | ||
| 60770 | 422 | text\<open>The wellordering theorem\<close> | 
| 13558 | 423 | theorem AC_well_ord: "\<exists>r. well_ord(S,r)" | 
| 13134 | 424 | apply (rule AC_Pi_Pow [THEN exE]) | 
| 13269 | 425 | apply (rule Zermelo_next_exists [THEN bexE], assumption) | 
| 13134 | 426 | apply (rule exI) | 
| 427 | apply (rule well_ord_rvimage) | |
| 428 | apply (erule_tac [2] well_ord_TFin) | |
| 13269 | 429 | apply (rule choice_imp_injection [THEN inj_weaken_type], blast+) | 
| 13134 | 430 | done | 
| 13558 | 431 | |
| 27704 | 432 | |
| 60770 | 433 | subsection \<open>Zorn's Lemma for Partial Orders\<close> | 
| 27704 | 434 | |
| 60770 | 435 | text \<open>Reimported from HOL by Clemens Ballarin.\<close> | 
| 27704 | 436 | |
| 437 | ||
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 438 | definition Chain :: "i \<Rightarrow> i" where | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 439 |   "Chain(r) = {A \<in> Pow(field(r)). \<forall>a\<in>A. \<forall>b\<in>A. \<langle>a, b\<rangle> \<in> r | \<langle>b, a\<rangle> \<in> r}"
 | 
| 27704 | 440 | |
| 441 | lemma mono_Chain: | |
| 76213 
e44d86131648
Removal of obsolete ASCII syntax
 paulson <lp15@cam.ac.uk> parents: 
69593diff
changeset | 442 | "r \<subseteq> s \<Longrightarrow> Chain(r) \<subseteq> Chain(s)" | 
| 27704 | 443 | unfolding Chain_def | 
| 444 | by blast | |
| 445 | ||
| 446 | theorem Zorn_po: | |
| 447 | assumes po: "Partial_order(r)" | |
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 448 | and u: "\<forall>C\<in>Chain(r). \<exists>u\<in>field(r). \<forall>a\<in>C. \<langle>a, u\<rangle> \<in> r" | 
| 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 449 | shows "\<exists>m\<in>field(r). \<forall>a\<in>field(r). \<langle>m, a\<rangle> \<in> r \<longrightarrow> a = m" | 
| 27704 | 450 | proof - | 
| 451 | have "Preorder(r)" using po by (simp add: partial_order_on_def) | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
61980diff
changeset | 452 | \<comment> \<open>Mirror r in the set of subsets below (wrt r) elements of A (?).\<close> | 
| 46820 | 453 |   let ?B = "\<lambda>x\<in>field(r). r -`` {x}" let ?S = "?B `` field(r)"
 | 
| 454 | have "\<forall>C\<in>chain(?S). \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U" | |
| 27704 | 455 | proof (clarsimp simp: chain_def Subset_rel_def bex_image_simp) | 
| 456 | fix C | |
| 46820 | 457 | assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B | B \<subseteq> A" | 
| 458 |     let ?A = "{x \<in> field(r). \<exists>M\<in>C. M = ?B`x}"
 | |
| 27704 | 459 | have "C = ?B `` ?A" using 1 | 
| 460 | apply (auto simp: image_def) | |
| 461 | apply rule | |
| 462 | apply rule | |
| 463 | apply (drule subsetD) apply assumption | |
| 464 | apply (erule CollectE) | |
| 465 | apply rule apply assumption | |
| 466 | apply (erule bexE) | |
| 467 | apply rule prefer 2 apply assumption | |
| 468 | apply rule | |
| 469 | apply (erule lamE) apply simp | |
| 470 | apply assumption | |
| 471 | ||
| 59788 | 472 | apply (thin_tac "C \<subseteq> X" for X) | 
| 27704 | 473 | apply (fast elim: lamE) | 
| 474 | done | |
| 46820 | 475 | have "?A \<in> Chain(r)" | 
| 27704 | 476 | proof (simp add: Chain_def subsetI, intro conjI ballI impI) | 
| 477 | fix a b | |
| 46820 | 478 |       assume "a \<in> field(r)" "r -`` {a} \<in> C" "b \<in> field(r)" "r -`` {b} \<in> C"
 | 
| 27704 | 479 |       hence "r -`` {a} \<subseteq> r -`` {b} | r -`` {b} \<subseteq> r -`` {a}" using 2 by auto
 | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 480 | then show "\<langle>a, b\<rangle> \<in> r | \<langle>b, a\<rangle> \<in> r" | 
| 60770 | 481 | using \<open>Preorder(r)\<close> \<open>a \<in> field(r)\<close> \<open>b \<in> field(r)\<close> | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 482 | by (simp add: subset_vimage1_vimage1_iff) | 
| 27704 | 483 | qed | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 484 | then obtain u where uA: "u \<in> field(r)" "\<forall>a\<in>?A. \<langle>a, u\<rangle> \<in> r" | 
| 27704 | 485 | using u | 
| 486 | apply auto | |
| 487 | apply (drule bspec) apply assumption | |
| 488 | apply auto | |
| 489 | done | |
| 46820 | 490 |     have "\<forall>A\<in>C. A \<subseteq> r -`` {u}"
 | 
| 27704 | 491 | proof (auto intro!: vimageI) | 
| 492 | fix a B | |
| 46820 | 493 | assume aB: "B \<in> C" "a \<in> B" | 
| 494 |       with 1 obtain x where "x \<in> field(r)" "B = r -`` {x}"
 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 495 | apply - | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 496 | apply (drule subsetD) apply assumption | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 497 | apply (erule imageE) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 498 | apply (erule lamE) | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 499 | apply simp | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 500 | done | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 501 | then show "\<langle>a, u\<rangle> \<in> r" using uA aB \<open>Preorder(r)\<close> | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
27704diff
changeset | 502 | by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+ | 
| 27704 | 503 | qed | 
| 46820 | 504 |     then show "\<exists>U\<in>field(r). \<forall>A\<in>C. A \<subseteq> r -`` {U}"
 | 
| 60770 | 505 | using \<open>u \<in> field(r)\<close> .. | 
| 27704 | 506 | qed | 
| 507 | from Zorn2 [OF this] | |
| 46820 | 508 |   obtain m B where "m \<in> field(r)" "B = r -`` {m}"
 | 
| 509 |     "\<forall>x\<in>field(r). B \<subseteq> r -`` {x} \<longrightarrow> B = r -`` {x}"
 | |
| 27704 | 510 | by (auto elim!: lamE simp: ball_image_simp) | 
| 76215 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 paulson <lp15@cam.ac.uk> parents: 
76214diff
changeset | 511 | then have "\<forall>a\<in>field(r). \<langle>m, a\<rangle> \<in> r \<longrightarrow> a = m" | 
| 60770 | 512 | using po \<open>Preorder(r)\<close> \<open>m \<in> field(r)\<close> | 
| 27704 | 513 | by (auto simp: subset_vimage1_vimage1_iff Partial_order_eq_vimage1_vimage1_iff) | 
| 60770 | 514 | then show ?thesis using \<open>m \<in> field(r)\<close> by blast | 
| 27704 | 515 | qed | 
| 516 | ||
| 516 | 517 | end |