src/HOL/Zorn.thy
changeset 63572 c0cbfd2b5a45
parent 63172 d4f459eb7ed0
child 67399 eab6ce8368fa
equal deleted inserted replaced
63571:aee0d92995b6 63572:c0cbfd2b5a45
     1 (*  Title:      HOL/Zorn.thy
     1 (*  Title:       HOL/Zorn.thy
     2     Author:     Jacques D. Fleuriot
     2     Author:      Jacques D. Fleuriot
     3     Author:     Tobias Nipkow, TUM
     3     Author:      Tobias Nipkow, TUM
     4     Author:     Christian Sternagel, JAIST
     4     Author:      Christian Sternagel, JAIST
     5 
     5 
     6 Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
     6 Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
     7 The well-ordering theorem.
     7 The well-ordering theorem.
     8 *)
     8 *)
     9 
     9 
    10 section \<open>Zorn's Lemma\<close>
    10 section \<open>Zorn's Lemma\<close>
    11 
    11 
    12 theory Zorn
    12 theory Zorn
    13 imports Order_Relation Hilbert_Choice
    13   imports Order_Relation Hilbert_Choice
    14 begin
    14 begin
    15 
    15 
    16 subsection \<open>Zorn's Lemma for the Subset Relation\<close>
    16 subsection \<open>Zorn's Lemma for the Subset Relation\<close>
    17 
    17 
    18 subsubsection \<open>Results that do not require an order\<close>
    18 subsubsection \<open>Results that do not require an order\<close>
    19 
    19 
    20 text \<open>Let \<open>P\<close> be a binary predicate on the set \<open>A\<close>.\<close>
    20 text \<open>Let \<open>P\<close> be a binary predicate on the set \<open>A\<close>.\<close>
    21 locale pred_on =
    21 locale pred_on =
    22   fixes A :: "'a set"
    22   fixes A :: "'a set"
    23     and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubset>" 50)
    23     and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infix "\<sqsubset>" 50)
    24 begin
    24 begin
    25 
    25 
    26 abbreviation Peq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) where
    26 abbreviation Peq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  (infix "\<sqsubseteq>" 50)
    27   "x \<sqsubseteq> y \<equiv> P\<^sup>=\<^sup>= x y"
    27   where "x \<sqsubseteq> y \<equiv> P\<^sup>=\<^sup>= x y"
    28 
    28 
    29 text \<open>A chain is a totally ordered subset of @{term A}.\<close>
    29 text \<open>A chain is a totally ordered subset of \<open>A\<close>.\<close>
    30 definition chain :: "'a set \<Rightarrow> bool" where
    30 definition chain :: "'a set \<Rightarrow> bool"
    31   "chain C \<longleftrightarrow> C \<subseteq> A \<and> (\<forall>x\<in>C. \<forall>y\<in>C. x \<sqsubseteq> y \<or> y \<sqsubseteq> x)"
    31   where "chain C \<longleftrightarrow> C \<subseteq> A \<and> (\<forall>x\<in>C. \<forall>y\<in>C. x \<sqsubseteq> y \<or> y \<sqsubseteq> x)"
    32 
    32 
    33 text \<open>We call a chain that is a proper superset of some set @{term X},
    33 text \<open>
    34 but not necessarily a chain itself, a superchain of @{term X}.\<close>
    34   We call a chain that is a proper superset of some set \<open>X\<close>,
    35 abbreviation superchain :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "<c" 50) where
    35   but not necessarily a chain itself, a superchain of \<open>X\<close>.
    36   "X <c C \<equiv> chain C \<and> X \<subset> C"
    36 \<close>
       
    37 abbreviation superchain :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"  (infix "<c" 50)
       
    38   where "X <c C \<equiv> chain C \<and> X \<subset> C"
    37 
    39 
    38 text \<open>A maximal chain is a chain that does not have a superchain.\<close>
    40 text \<open>A maximal chain is a chain that does not have a superchain.\<close>
    39 definition maxchain :: "'a set \<Rightarrow> bool" where
    41 definition maxchain :: "'a set \<Rightarrow> bool"
    40   "maxchain C \<longleftrightarrow> chain C \<and> \<not> (\<exists>S. C <c S)"
    42   where "maxchain C \<longleftrightarrow> chain C \<and> (\<nexists>S. C <c S)"
    41 
    43 
    42 text \<open>We define the successor of a set to be an arbitrary
    44 text \<open>
    43 superchain, if such exists, or the set itself, otherwise.\<close>
    45   We define the successor of a set to be an arbitrary
    44 definition suc :: "'a set \<Rightarrow> 'a set" where
    46   superchain, if such exists, or the set itself, otherwise.
    45   "suc C = (if \<not> chain C \<or> maxchain C then C else (SOME D. C <c D))"
    47 \<close>
    46 
    48 definition suc :: "'a set \<Rightarrow> 'a set"
    47 lemma chainI [Pure.intro?]:
    49   where "suc C = (if \<not> chain C \<or> maxchain C then C else (SOME D. C <c D))"
    48   "\<lbrakk>C \<subseteq> A; \<And>x y. \<lbrakk>x \<in> C; y \<in> C\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x\<rbrakk> \<Longrightarrow> chain C"
    50 
       
    51 lemma chainI [Pure.intro?]: "C \<subseteq> A \<Longrightarrow> (\<And>x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x) \<Longrightarrow> chain C"
    49   unfolding chain_def by blast
    52   unfolding chain_def by blast
    50 
    53 
    51 lemma chain_total:
    54 lemma chain_total: "chain C \<Longrightarrow> x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
    52   "chain C \<Longrightarrow> x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
       
    53   by (simp add: chain_def)
    55   by (simp add: chain_def)
    54 
    56 
    55 lemma not_chain_suc [simp]: "\<not> chain X \<Longrightarrow> suc X = X"
    57 lemma not_chain_suc [simp]: "\<not> chain X \<Longrightarrow> suc X = X"
    56   by (simp add: suc_def)
    58   by (simp add: suc_def)
    57 
    59 
    62   by (auto simp: suc_def maxchain_def intro: someI2)
    64   by (auto simp: suc_def maxchain_def intro: someI2)
    63 
    65 
    64 lemma chain_empty [simp]: "chain {}"
    66 lemma chain_empty [simp]: "chain {}"
    65   by (auto simp: chain_def)
    67   by (auto simp: chain_def)
    66 
    68 
    67 lemma not_maxchain_Some:
    69 lemma not_maxchain_Some: "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> C <c (SOME D. C <c D)"
    68   "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> C <c (SOME D. C <c D)"
       
    69   by (rule someI_ex) (auto simp: maxchain_def)
    70   by (rule someI_ex) (auto simp: maxchain_def)
    70 
    71 
    71 lemma suc_not_equals:
    72 lemma suc_not_equals: "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> suc C \<noteq> C"
    72   "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> suc C \<noteq> C"
       
    73   using not_maxchain_Some by (auto simp: suc_def)
    73   using not_maxchain_Some by (auto simp: suc_def)
    74 
    74 
    75 lemma subset_suc:
    75 lemma subset_suc:
    76   assumes "X \<subseteq> Y" shows "X \<subseteq> suc Y"
    76   assumes "X \<subseteq> Y"
       
    77   shows "X \<subseteq> suc Y"
    77   using assms by (rule subset_trans) (rule suc_subset)
    78   using assms by (rule subset_trans) (rule suc_subset)
    78 
    79 
    79 text \<open>We build a set @{term \<C>} that is closed under applications
    80 text \<open>
    80 of @{term suc} and contains the union of all its subsets.\<close>
    81   We build a set @{term \<C>} that is closed under applications
    81 inductive_set suc_Union_closed ("\<C>") where
    82   of @{term suc} and contains the union of all its subsets.
    82   suc: "X \<in> \<C> \<Longrightarrow> suc X \<in> \<C>" |
    83 \<close>
    83   Union [unfolded Pow_iff]: "X \<in> Pow \<C> \<Longrightarrow> \<Union>X \<in> \<C>"
    84 inductive_set suc_Union_closed ("\<C>")
    84 
    85   where
    85 text \<open>Since the empty set as well as the set itself is a subset of
    86     suc: "X \<in> \<C> \<Longrightarrow> suc X \<in> \<C>"
    86 every set, @{term \<C>} contains at least @{term "{} \<in> \<C>"} and
    87   | Union [unfolded Pow_iff]: "X \<in> Pow \<C> \<Longrightarrow> \<Union>X \<in> \<C>"
    87 @{term "\<Union>\<C> \<in> \<C>"}.\<close>
    88 
    88 lemma
    89 text \<open>
    89   suc_Union_closed_empty: "{} \<in> \<C>" and
    90   Since the empty set as well as the set itself is a subset of
    90   suc_Union_closed_Union: "\<Union>\<C> \<in> \<C>"
    91   every set, @{term \<C>} contains at least @{term "{} \<in> \<C>"} and
    91   using Union [of "{}"] and Union [of "\<C>"] by simp+
    92   @{term "\<Union>\<C> \<in> \<C>"}.
       
    93 \<close>
       
    94 lemma suc_Union_closed_empty: "{} \<in> \<C>"
       
    95   and suc_Union_closed_Union: "\<Union>\<C> \<in> \<C>"
       
    96   using Union [of "{}"] and Union [of "\<C>"] by simp_all
       
    97 
    92 text \<open>Thus closure under @{term suc} will hit a maximal chain
    98 text \<open>Thus closure under @{term suc} will hit a maximal chain
    93 eventually, as is shown below.\<close>
    99   eventually, as is shown below.\<close>
    94 
   100 
    95 lemma suc_Union_closed_induct [consumes 1, case_names suc Union,
   101 lemma suc_Union_closed_induct [consumes 1, case_names suc Union, induct pred: suc_Union_closed]:
    96   induct pred: suc_Union_closed]:
       
    97   assumes "X \<in> \<C>"
   102   assumes "X \<in> \<C>"
    98     and "\<And>X. \<lbrakk>X \<in> \<C>; Q X\<rbrakk> \<Longrightarrow> Q (suc X)"
   103     and "\<And>X. X \<in> \<C> \<Longrightarrow> Q X \<Longrightarrow> Q (suc X)"
    99     and "\<And>X. \<lbrakk>X \<subseteq> \<C>; \<forall>x\<in>X. Q x\<rbrakk> \<Longrightarrow> Q (\<Union>X)"
   104     and "\<And>X. X \<subseteq> \<C> \<Longrightarrow> \<forall>x\<in>X. Q x \<Longrightarrow> Q (\<Union>X)"
   100   shows "Q X"
   105   shows "Q X"
   101   using assms by (induct) blast+
   106   using assms by induct blast+
   102 
   107 
   103 lemma suc_Union_closed_cases [consumes 1, case_names suc Union,
   108 lemma suc_Union_closed_cases [consumes 1, case_names suc Union, cases pred: suc_Union_closed]:
   104   cases pred: suc_Union_closed]:
       
   105   assumes "X \<in> \<C>"
   109   assumes "X \<in> \<C>"
   106     and "\<And>Y. \<lbrakk>X = suc Y; Y \<in> \<C>\<rbrakk> \<Longrightarrow> Q"
   110     and "\<And>Y. X = suc Y \<Longrightarrow> Y \<in> \<C> \<Longrightarrow> Q"
   107     and "\<And>Y. \<lbrakk>X = \<Union>Y; Y \<subseteq> \<C>\<rbrakk> \<Longrightarrow> Q"
   111     and "\<And>Y. X = \<Union>Y \<Longrightarrow> Y \<subseteq> \<C> \<Longrightarrow> Q"
   108   shows "Q"
   112   shows "Q"
   109   using assms by (cases) simp+
   113   using assms by cases simp_all
   110 
   114 
   111 text \<open>On chains, @{term suc} yields a chain.\<close>
   115 text \<open>On chains, @{term suc} yields a chain.\<close>
   112 lemma chain_suc:
   116 lemma chain_suc:
   113   assumes "chain X" shows "chain (suc X)"
   117   assumes "chain X"
       
   118   shows "chain (suc X)"
   114   using assms
   119   using assms
   115   by (cases "\<not> chain X \<or> maxchain X")
   120   by (cases "\<not> chain X \<or> maxchain X") (force simp: suc_def dest: not_maxchain_Some)+
   116      (force simp: suc_def dest: not_maxchain_Some)+
       
   117 
   121 
   118 lemma chain_sucD:
   122 lemma chain_sucD:
   119   assumes "chain X" shows "suc X \<subseteq> A \<and> chain (suc X)"
   123   assumes "chain X"
       
   124   shows "suc X \<subseteq> A \<and> chain (suc X)"
   120 proof -
   125 proof -
   121   from \<open>chain X\<close> have *: "chain (suc X)" by (rule chain_suc)
   126   from \<open>chain X\<close> have *: "chain (suc X)"
   122   then have "suc X \<subseteq> A" unfolding chain_def by blast
   127     by (rule chain_suc)
       
   128   then have "suc X \<subseteq> A"
       
   129     unfolding chain_def by blast
   123   with * show ?thesis by blast
   130   with * show ?thesis by blast
   124 qed
   131 qed
   125 
   132 
   126 lemma suc_Union_closed_total':
   133 lemma suc_Union_closed_total':
   127   assumes "X \<in> \<C>" and "Y \<in> \<C>"
   134   assumes "X \<in> \<C>" and "Y \<in> \<C>"
   128     and *: "\<And>Z. Z \<in> \<C> \<Longrightarrow> Z \<subseteq> Y \<Longrightarrow> Z = Y \<or> suc Z \<subseteq> Y"
   135     and *: "\<And>Z. Z \<in> \<C> \<Longrightarrow> Z \<subseteq> Y \<Longrightarrow> Z = Y \<or> suc Z \<subseteq> Y"
   129   shows "X \<subseteq> Y \<or> suc Y \<subseteq> X"
   136   shows "X \<subseteq> Y \<or> suc Y \<subseteq> X"
   130   using \<open>X \<in> \<C>\<close>
   137   using \<open>X \<in> \<C>\<close>
   131 proof (induct)
   138 proof induct
   132   case (suc X)
   139   case (suc X)
   133   with * show ?case by (blast del: subsetI intro: subset_suc)
   140   with * show ?case by (blast del: subsetI intro: subset_suc)
   134 qed blast
   141 next
       
   142   case Union
       
   143   then show ?case by blast
       
   144 qed
   135 
   145 
   136 lemma suc_Union_closed_subsetD:
   146 lemma suc_Union_closed_subsetD:
   137   assumes "Y \<subseteq> X" and "X \<in> \<C>" and "Y \<in> \<C>"
   147   assumes "Y \<subseteq> X" and "X \<in> \<C>" and "Y \<in> \<C>"
   138   shows "X = Y \<or> suc Y \<subseteq> X"
   148   shows "X = Y \<or> suc Y \<subseteq> X"
   139   using assms(2-, 1)
   149   using assms(2,3,1)
   140 proof (induct arbitrary: Y)
   150 proof (induct arbitrary: Y)
   141   case (suc X)
   151   case (suc X)
   142   note * = \<open>\<And>Y. \<lbrakk>Y \<in> \<C>; Y \<subseteq> X\<rbrakk> \<Longrightarrow> X = Y \<or> suc Y \<subseteq> X\<close>
   152   note * = \<open>\<And>Y. Y \<in> \<C> \<Longrightarrow> Y \<subseteq> X \<Longrightarrow> X = Y \<or> suc Y \<subseteq> X\<close>
   143   with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>X \<in> \<C>\<close>]
   153   with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>X \<in> \<C>\<close>]
   144     have "Y \<subseteq> X \<or> suc X \<subseteq> Y" by blast
   154   have "Y \<subseteq> X \<or> suc X \<subseteq> Y" by blast
   145   then show ?case
   155   then show ?case
   146   proof
   156   proof
   147     assume "Y \<subseteq> X"
   157     assume "Y \<subseteq> X"
   148     with * and \<open>Y \<in> \<C>\<close> have "X = Y \<or> suc Y \<subseteq> X" by blast
   158     with * and \<open>Y \<in> \<C>\<close> have "X = Y \<or> suc Y \<subseteq> X" by blast
   149     then show ?thesis
   159     then show ?thesis
   150     proof
   160     proof
   151       assume "X = Y" then show ?thesis by simp
   161       assume "X = Y"
       
   162       then show ?thesis by simp
   152     next
   163     next
   153       assume "suc Y \<subseteq> X"
   164       assume "suc Y \<subseteq> X"
   154       then have "suc Y \<subseteq> suc X" by (rule subset_suc)
   165       then have "suc Y \<subseteq> suc X" by (rule subset_suc)
   155       then show ?thesis by simp
   166       then show ?thesis by simp
   156     qed
   167     qed
   162   case (Union X)
   173   case (Union X)
   163   show ?case
   174   show ?case
   164   proof (rule ccontr)
   175   proof (rule ccontr)
   165     assume "\<not> ?thesis"
   176     assume "\<not> ?thesis"
   166     with \<open>Y \<subseteq> \<Union>X\<close> obtain x y z
   177     with \<open>Y \<subseteq> \<Union>X\<close> obtain x y z
   167     where "\<not> suc Y \<subseteq> \<Union>X"
   178       where "\<not> suc Y \<subseteq> \<Union>X"
   168       and "x \<in> X" and "y \<in> x" and "y \<notin> Y"
   179         and "x \<in> X" and "y \<in> x" and "y \<notin> Y"
   169       and "z \<in> suc Y" and "\<forall>x\<in>X. z \<notin> x" by blast
   180         and "z \<in> suc Y" and "\<forall>x\<in>X. z \<notin> x" by blast
   170     with \<open>X \<subseteq> \<C>\<close> have "x \<in> \<C>" by blast
   181     with \<open>X \<subseteq> \<C>\<close> have "x \<in> \<C>" by blast
   171     from Union and \<open>x \<in> X\<close>
   182     from Union and \<open>x \<in> X\<close> have *: "\<And>y. y \<in> \<C> \<Longrightarrow> y \<subseteq> x \<Longrightarrow> x = y \<or> suc y \<subseteq> x"
   172       have *: "\<And>y. \<lbrakk>y \<in> \<C>; y \<subseteq> x\<rbrakk> \<Longrightarrow> x = y \<or> suc y \<subseteq> x" by blast
   183       by blast
   173     with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>x \<in> \<C>\<close>]
   184     with suc_Union_closed_total' [OF \<open>Y \<in> \<C>\<close> \<open>x \<in> \<C>\<close>] have "Y \<subseteq> x \<or> suc x \<subseteq> Y"
   174       have "Y \<subseteq> x \<or> suc x \<subseteq> Y" by blast
   185       by blast
   175     then show False
   186     then show False
   176     proof
   187     proof
   177       assume "Y \<subseteq> x"
   188       assume "Y \<subseteq> x"
   178       with * [OF \<open>Y \<in> \<C>\<close>] have "x = Y \<or> suc Y \<subseteq> x" by blast
   189       with * [OF \<open>Y \<in> \<C>\<close>] have "x = Y \<or> suc Y \<subseteq> x" by blast
   179       then show False
   190       then show False
   180       proof
   191       proof
   181         assume "x = Y" with \<open>y \<in> x\<close> and \<open>y \<notin> Y\<close> show False by blast
   192         assume "x = Y"
       
   193         with \<open>y \<in> x\<close> and \<open>y \<notin> Y\<close> show False by blast
   182       next
   194       next
   183         assume "suc Y \<subseteq> x"
   195         assume "suc Y \<subseteq> x"
   184         with \<open>x \<in> X\<close> have "suc Y \<subseteq> \<Union>X" by blast
   196         with \<open>x \<in> X\<close> have "suc Y \<subseteq> \<Union>X" by blast
   185         with \<open>\<not> suc Y \<subseteq> \<Union>X\<close> show False by contradiction
   197         with \<open>\<not> suc Y \<subseteq> \<Union>X\<close> show False by contradiction
   186       qed
   198       qed
   197   assumes "X \<in> \<C>" and "Y \<in> \<C>"
   209   assumes "X \<in> \<C>" and "Y \<in> \<C>"
   198   shows "X \<subseteq> Y \<or> Y \<subseteq> X"
   210   shows "X \<subseteq> Y \<or> Y \<subseteq> X"
   199 proof (cases "\<forall>Z\<in>\<C>. Z \<subseteq> Y \<longrightarrow> Z = Y \<or> suc Z \<subseteq> Y")
   211 proof (cases "\<forall>Z\<in>\<C>. Z \<subseteq> Y \<longrightarrow> Z = Y \<or> suc Z \<subseteq> Y")
   200   case True
   212   case True
   201   with suc_Union_closed_total' [OF assms]
   213   with suc_Union_closed_total' [OF assms]
   202     have "X \<subseteq> Y \<or> suc Y \<subseteq> X" by blast
   214   have "X \<subseteq> Y \<or> suc Y \<subseteq> X" by blast
   203   then show ?thesis using suc_subset [of Y] by blast
   215   with suc_subset [of Y] show ?thesis by blast
   204 next
   216 next
   205   case False
   217   case False
   206   then obtain Z
   218   then obtain Z where "Z \<in> \<C>" and "Z \<subseteq> Y" and "Z \<noteq> Y" and "\<not> suc Z \<subseteq> Y"
   207     where "Z \<in> \<C>" and "Z \<subseteq> Y" and "Z \<noteq> Y" and "\<not> suc Z \<subseteq> Y" by blast
   219     by blast
   208   with suc_Union_closed_subsetD and \<open>Y \<in> \<C>\<close> show ?thesis by blast
   220   with suc_Union_closed_subsetD and \<open>Y \<in> \<C>\<close> show ?thesis
       
   221     by blast
   209 qed
   222 qed
   210 
   223 
   211 text \<open>Once we hit a fixed point w.r.t. @{term suc}, all other elements
   224 text \<open>Once we hit a fixed point w.r.t. @{term suc}, all other elements
   212 of @{term \<C>} are subsets of this fixed point.\<close>
   225   of @{term \<C>} are subsets of this fixed point.\<close>
   213 lemma suc_Union_closed_suc:
   226 lemma suc_Union_closed_suc:
   214   assumes "X \<in> \<C>" and "Y \<in> \<C>" and "suc Y = Y"
   227   assumes "X \<in> \<C>" and "Y \<in> \<C>" and "suc Y = Y"
   215   shows "X \<subseteq> Y"
   228   shows "X \<subseteq> Y"
   216 using \<open>X \<in> \<C>\<close>
   229   using \<open>X \<in> \<C>\<close>
   217 proof (induct)
   230 proof induct
   218   case (suc X)
   231   case (suc X)
   219   with \<open>Y \<in> \<C>\<close> and suc_Union_closed_subsetD
   232   with \<open>Y \<in> \<C>\<close> and suc_Union_closed_subsetD have "X = Y \<or> suc X \<subseteq> Y"
   220     have "X = Y \<or> suc X \<subseteq> Y" by blast
   233     by blast
   221   then show ?case by (auto simp: \<open>suc Y = Y\<close>)
   234   then show ?case
   222 qed blast
   235     by (auto simp: \<open>suc Y = Y\<close>)
       
   236 next
       
   237   case Union
       
   238   then show ?case by blast
       
   239 qed
   223 
   240 
   224 lemma eq_suc_Union:
   241 lemma eq_suc_Union:
   225   assumes "X \<in> \<C>"
   242   assumes "X \<in> \<C>"
   226   shows "suc X = X \<longleftrightarrow> X = \<Union>\<C>"
   243   shows "suc X = X \<longleftrightarrow> X = \<Union>\<C>"
       
   244     (is "?lhs \<longleftrightarrow> ?rhs")
   227 proof
   245 proof
   228   assume "suc X = X"
   246   assume ?lhs
   229   with suc_Union_closed_suc [OF suc_Union_closed_Union \<open>X \<in> \<C>\<close>]
   247   then have "\<Union>\<C> \<subseteq> X"
   230     have "\<Union>\<C> \<subseteq> X" .
   248     by (rule suc_Union_closed_suc [OF suc_Union_closed_Union \<open>X \<in> \<C>\<close>])
   231   with \<open>X \<in> \<C>\<close> show "X = \<Union>\<C>" by blast
   249   with \<open>X \<in> \<C>\<close> show ?rhs
       
   250     by blast
   232 next
   251 next
   233   from \<open>X \<in> \<C>\<close> have "suc X \<in> \<C>" by (rule suc)
   252   from \<open>X \<in> \<C>\<close> have "suc X \<in> \<C>" by (rule suc)
   234   then have "suc X \<subseteq> \<Union>\<C>" by blast
   253   then have "suc X \<subseteq> \<Union>\<C>" by blast
   235   moreover assume "X = \<Union>\<C>"
   254   moreover assume ?rhs
   236   ultimately have "suc X \<subseteq> X" by simp
   255   ultimately have "suc X \<subseteq> X" by simp
   237   moreover have "X \<subseteq> suc X" by (rule suc_subset)
   256   moreover have "X \<subseteq> suc X" by (rule suc_subset)
   238   ultimately show "suc X = X" ..
   257   ultimately show ?lhs ..
   239 qed
   258 qed
   240 
   259 
   241 lemma suc_in_carrier:
   260 lemma suc_in_carrier:
   242   assumes "X \<subseteq> A"
   261   assumes "X \<subseteq> A"
   243   shows "suc X \<subseteq> A"
   262   shows "suc X \<subseteq> A"
   244   using assms
   263   using assms
   245   by (cases "\<not> chain X \<or> maxchain X")
   264   by (cases "\<not> chain X \<or> maxchain X") (auto dest: chain_sucD)
   246      (auto dest: chain_sucD)
       
   247 
   265 
   248 lemma suc_Union_closed_in_carrier:
   266 lemma suc_Union_closed_in_carrier:
   249   assumes "X \<in> \<C>"
   267   assumes "X \<in> \<C>"
   250   shows "X \<subseteq> A"
   268   shows "X \<subseteq> A"
   251   using assms
   269   using assms
   252   by (induct) (auto dest: suc_in_carrier)
   270   by induct (auto dest: suc_in_carrier)
   253 
   271 
   254 text \<open>All elements of @{term \<C>} are chains.\<close>
   272 text \<open>All elements of @{term \<C>} are chains.\<close>
   255 lemma suc_Union_closed_chain:
   273 lemma suc_Union_closed_chain:
   256   assumes "X \<in> \<C>"
   274   assumes "X \<in> \<C>"
   257   shows "chain X"
   275   shows "chain X"
   258 using assms
   276   using assms
   259 proof (induct)
   277 proof induct
   260   case (suc X) then show ?case using not_maxchain_Some by (simp add: suc_def)
   278   case (suc X)
       
   279   then show ?case
       
   280     using not_maxchain_Some by (simp add: suc_def)
   261 next
   281 next
   262   case (Union X)
   282   case (Union X)
   263   then have "\<Union>X \<subseteq> A" by (auto dest: suc_Union_closed_in_carrier)
   283   then have "\<Union>X \<subseteq> A"
       
   284     by (auto dest: suc_Union_closed_in_carrier)
   264   moreover have "\<forall>x\<in>\<Union>X. \<forall>y\<in>\<Union>X. x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   285   moreover have "\<forall>x\<in>\<Union>X. \<forall>y\<in>\<Union>X. x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   265   proof (intro ballI)
   286   proof (intro ballI)
   266     fix x y
   287     fix x y
   267     assume "x \<in> \<Union>X" and "y \<in> \<Union>X"
   288     assume "x \<in> \<Union>X" and "y \<in> \<Union>X"
   268     then obtain u v where "x \<in> u" and "u \<in> X" and "y \<in> v" and "v \<in> X" by blast
   289     then obtain u v where "x \<in> u" and "u \<in> X" and "y \<in> v" and "v \<in> X"
   269     with Union have "u \<in> \<C>" and "v \<in> \<C>" and "chain u" and "chain v" by blast+
   290       by blast
   270     with suc_Union_closed_total have "u \<subseteq> v \<or> v \<subseteq> u" by blast
   291     with Union have "u \<in> \<C>" and "v \<in> \<C>" and "chain u" and "chain v"
       
   292       by blast+
       
   293     with suc_Union_closed_total have "u \<subseteq> v \<or> v \<subseteq> u"
       
   294       by blast
   271     then show "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   295     then show "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   272     proof
   296     proof
   273       assume "u \<subseteq> v"
   297       assume "u \<subseteq> v"
   274       from \<open>chain v\<close> show ?thesis
   298       from \<open>chain v\<close> show ?thesis
   275       proof (rule chain_total)
   299       proof (rule chain_total)
   288   ultimately show ?case unfolding chain_def ..
   312   ultimately show ?case unfolding chain_def ..
   289 qed
   313 qed
   290 
   314 
   291 subsubsection \<open>Hausdorff's Maximum Principle\<close>
   315 subsubsection \<open>Hausdorff's Maximum Principle\<close>
   292 
   316 
   293 text \<open>There exists a maximal totally ordered subset of @{term A}. (Note that we do not
   317 text \<open>There exists a maximal totally ordered subset of \<open>A\<close>. (Note that we do not
   294 require @{term A} to be partially ordered.)\<close>
   318   require \<open>A\<close> to be partially ordered.)\<close>
   295 
   319 
   296 theorem Hausdorff: "\<exists>C. maxchain C"
   320 theorem Hausdorff: "\<exists>C. maxchain C"
   297 proof -
   321 proof -
   298   let ?M = "\<Union>\<C>"
   322   let ?M = "\<Union>\<C>"
   299   have "maxchain ?M"
   323   have "maxchain ?M"
   300   proof (rule ccontr)
   324   proof (rule ccontr)
   301     assume "\<not> maxchain ?M"
   325     assume "\<not> ?thesis"
   302     then have "suc ?M \<noteq> ?M"
   326     then have "suc ?M \<noteq> ?M"
   303       using suc_not_equals and
   327       using suc_not_equals and suc_Union_closed_chain [OF suc_Union_closed_Union] by simp
   304       suc_Union_closed_chain [OF suc_Union_closed_Union] by simp
       
   305     moreover have "suc ?M = ?M"
   328     moreover have "suc ?M = ?M"
   306       using eq_suc_Union [OF suc_Union_closed_Union] by simp
   329       using eq_suc_Union [OF suc_Union_closed_Union] by simp
   307     ultimately show False by contradiction
   330     ultimately show False by contradiction
   308   qed
   331   qed
   309   then show ?thesis by blast
   332   then show ?thesis by blast
   310 qed
   333 qed
   311 
   334 
   312 text \<open>Make notation @{term \<C>} available again.\<close>
   335 text \<open>Make notation @{term \<C>} available again.\<close>
   313 no_notation suc_Union_closed ("\<C>")
   336 no_notation suc_Union_closed  ("\<C>")
   314 
   337 
   315 lemma chain_extend:
   338 lemma chain_extend: "chain C \<Longrightarrow> z \<in> A \<Longrightarrow> \<forall>x\<in>C. x \<sqsubseteq> z \<Longrightarrow> chain ({z} \<union> C)"
   316   "chain C \<Longrightarrow> z \<in> A \<Longrightarrow> \<forall>x\<in>C. x \<sqsubseteq> z \<Longrightarrow> chain ({z} \<union> C)"
       
   317   unfolding chain_def by blast
   339   unfolding chain_def by blast
   318 
   340 
   319 lemma maxchain_imp_chain:
   341 lemma maxchain_imp_chain: "maxchain C \<Longrightarrow> chain C"
   320   "maxchain C \<Longrightarrow> chain C"
       
   321   by (simp add: maxchain_def)
   342   by (simp add: maxchain_def)
   322 
   343 
   323 end
   344 end
   324 
   345 
   325 text \<open>Hide constant @{const pred_on.suc_Union_closed}, which was just needed
   346 text \<open>Hide constant @{const pred_on.suc_Union_closed}, which was just needed
   326 for the proof of Hausforff's maximum principle.\<close>
   347   for the proof of Hausforff's maximum principle.\<close>
   327 hide_const pred_on.suc_Union_closed
   348 hide_const pred_on.suc_Union_closed
   328 
   349 
   329 lemma chain_mono:
   350 lemma chain_mono:
   330   assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> A; P x y\<rbrakk> \<Longrightarrow> Q x y"
   351   assumes "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> P x y \<Longrightarrow> Q x y"
   331     and "pred_on.chain A P C"
   352     and "pred_on.chain A P C"
   332   shows "pred_on.chain A Q C"
   353   shows "pred_on.chain A Q C"
   333   using assms unfolding pred_on.chain_def by blast
   354   using assms unfolding pred_on.chain_def by blast
   334 
   355 
       
   356 
   335 subsubsection \<open>Results for the proper subset relation\<close>
   357 subsubsection \<open>Results for the proper subset relation\<close>
   336 
   358 
   337 interpretation subset: pred_on "A" "op \<subset>" for A .
   359 interpretation subset: pred_on "A" "op \<subset>" for A .
   338 
   360 
   339 lemma subset_maxchain_max:
   361 lemma subset_maxchain_max:
   340   assumes "subset.maxchain A C" and "X \<in> A" and "\<Union>C \<subseteq> X"
   362   assumes "subset.maxchain A C"
       
   363     and "X \<in> A"
       
   364     and "\<Union>C \<subseteq> X"
   341   shows "\<Union>C = X"
   365   shows "\<Union>C = X"
   342 proof (rule ccontr)
   366 proof (rule ccontr)
   343   let ?C = "{X} \<union> C"
   367   let ?C = "{X} \<union> C"
   344   from \<open>subset.maxchain A C\<close> have "subset.chain A C"
   368   from \<open>subset.maxchain A C\<close> have "subset.chain A C"
   345     and *: "\<And>S. subset.chain A S \<Longrightarrow> \<not> C \<subset> S"
   369     and *: "\<And>S. subset.chain A S \<Longrightarrow> \<not> C \<subset> S"
   350   moreover assume **: "\<Union>C \<noteq> X"
   374   moreover assume **: "\<Union>C \<noteq> X"
   351   moreover from ** have "C \<subset> ?C" using \<open>\<Union>C \<subseteq> X\<close> by auto
   375   moreover from ** have "C \<subset> ?C" using \<open>\<Union>C \<subseteq> X\<close> by auto
   352   ultimately show False using * by blast
   376   ultimately show False using * by blast
   353 qed
   377 qed
   354 
   378 
       
   379 
   355 subsubsection \<open>Zorn's lemma\<close>
   380 subsubsection \<open>Zorn's lemma\<close>
   356 
   381 
   357 text \<open>If every chain has an upper bound, then there is a maximal set.\<close>
   382 text \<open>If every chain has an upper bound, then there is a maximal set.\<close>
   358 lemma subset_Zorn:
   383 lemma subset_Zorn:
   359   assumes "\<And>C. subset.chain A C \<Longrightarrow> \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U"
   384   assumes "\<And>C. subset.chain A C \<Longrightarrow> \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U"
   360   shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   385   shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   361 proof -
   386 proof -
   362   from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
   387   from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
   363   then have "subset.chain A M" by (rule subset.maxchain_imp_chain)
   388   then have "subset.chain A M"
   364   with assms obtain Y where "Y \<in> A" and "\<forall>X\<in>M. X \<subseteq> Y" by blast
   389     by (rule subset.maxchain_imp_chain)
       
   390   with assms obtain Y where "Y \<in> A" and "\<forall>X\<in>M. X \<subseteq> Y"
       
   391     by blast
   365   moreover have "\<forall>X\<in>A. Y \<subseteq> X \<longrightarrow> Y = X"
   392   moreover have "\<forall>X\<in>A. Y \<subseteq> X \<longrightarrow> Y = X"
   366   proof (intro ballI impI)
   393   proof (intro ballI impI)
   367     fix X
   394     fix X
   368     assume "X \<in> A" and "Y \<subseteq> X"
   395     assume "X \<in> A" and "Y \<subseteq> X"
   369     show "Y = X"
   396     show "Y = X"
   370     proof (rule ccontr)
   397     proof (rule ccontr)
   371       assume "Y \<noteq> X"
   398       assume "\<not> ?thesis"
   372       with \<open>Y \<subseteq> X\<close> have "\<not> X \<subseteq> Y" by blast
   399       with \<open>Y \<subseteq> X\<close> have "\<not> X \<subseteq> Y" by blast
   373       from subset.chain_extend [OF \<open>subset.chain A M\<close> \<open>X \<in> A\<close>] and \<open>\<forall>X\<in>M. X \<subseteq> Y\<close>
   400       from subset.chain_extend [OF \<open>subset.chain A M\<close> \<open>X \<in> A\<close>] and \<open>\<forall>X\<in>M. X \<subseteq> Y\<close>
   374         have "subset.chain A ({X} \<union> M)" using \<open>Y \<subseteq> X\<close> by auto
   401       have "subset.chain A ({X} \<union> M)"
   375       moreover have "M \<subset> {X} \<union> M" using \<open>\<forall>X\<in>M. X \<subseteq> Y\<close> and \<open>\<not> X \<subseteq> Y\<close> by auto
   402         using \<open>Y \<subseteq> X\<close> by auto
       
   403       moreover have "M \<subset> {X} \<union> M"
       
   404         using \<open>\<forall>X\<in>M. X \<subseteq> Y\<close> and \<open>\<not> X \<subseteq> Y\<close> by auto
   376       ultimately show False
   405       ultimately show False
   377         using \<open>subset.maxchain A M\<close> by (auto simp: subset.maxchain_def)
   406         using \<open>subset.maxchain A M\<close> by (auto simp: subset.maxchain_def)
   378     qed
   407     qed
   379   qed
   408   qed
   380   ultimately show ?thesis by blast
   409   ultimately show ?thesis by blast
   381 qed
   410 qed
   382 
   411 
   383 text\<open>Alternative version of Zorn's lemma for the subset relation.\<close>
   412 text \<open>Alternative version of Zorn's lemma for the subset relation.\<close>
   384 lemma subset_Zorn':
   413 lemma subset_Zorn':
   385   assumes "\<And>C. subset.chain A C \<Longrightarrow> \<Union>C \<in> A"
   414   assumes "\<And>C. subset.chain A C \<Longrightarrow> \<Union>C \<in> A"
   386   shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   415   shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   387 proof -
   416 proof -
   388   from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
   417   from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
   389   then have "subset.chain A M" by (rule subset.maxchain_imp_chain)
   418   then have "subset.chain A M"
       
   419     by (rule subset.maxchain_imp_chain)
   390   with assms have "\<Union>M \<in> A" .
   420   with assms have "\<Union>M \<in> A" .
   391   moreover have "\<forall>Z\<in>A. \<Union>M \<subseteq> Z \<longrightarrow> \<Union>M = Z"
   421   moreover have "\<forall>Z\<in>A. \<Union>M \<subseteq> Z \<longrightarrow> \<Union>M = Z"
   392   proof (intro ballI impI)
   422   proof (intro ballI impI)
   393     fix Z
   423     fix Z
   394     assume "Z \<in> A" and "\<Union>M \<subseteq> Z"
   424     assume "Z \<in> A" and "\<Union>M \<subseteq> Z"
   401 
   431 
   402 subsection \<open>Zorn's Lemma for Partial Orders\<close>
   432 subsection \<open>Zorn's Lemma for Partial Orders\<close>
   403 
   433 
   404 text \<open>Relate old to new definitions.\<close>
   434 text \<open>Relate old to new definitions.\<close>
   405 
   435 
   406 (* Define globally? In Set.thy? *)
   436 definition chain_subset :: "'a set set \<Rightarrow> bool"  ("chain\<^sub>\<subseteq>")  (* Define globally? In Set.thy? *)
   407 definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^sub>\<subseteq>") where
   437   where "chain\<^sub>\<subseteq> C \<longleftrightarrow> (\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A)"
   408   "chain\<^sub>\<subseteq> C \<longleftrightarrow> (\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A)"
   438 
   409 
   439 definition chains :: "'a set set \<Rightarrow> 'a set set set"
   410 definition chains :: "'a set set \<Rightarrow> 'a set set set" where
   440   where "chains A = {C. C \<subseteq> A \<and> chain\<^sub>\<subseteq> C}"
   411   "chains A = {C. C \<subseteq> A \<and> chain\<^sub>\<subseteq> C}"
   441 
   412 
   442 definition Chains :: "('a \<times> 'a) set \<Rightarrow> 'a set set"  (* Define globally? In Relation.thy? *)
   413 (* Define globally? In Relation.thy? *)
   443   where "Chains r = {C. \<forall>a\<in>C. \<forall>b\<in>C. (a, b) \<in> r \<or> (b, a) \<in> r}"
   414 definition Chains :: "('a \<times> 'a) set \<Rightarrow> 'a set set" where
   444 
   415   "Chains r = {C. \<forall>a\<in>C. \<forall>b\<in>C. (a, b) \<in> r \<or> (b, a) \<in> r}"
   445 lemma chains_extend: "c \<in> chains S \<Longrightarrow> z \<in> S \<Longrightarrow> \<forall>x \<in> c. x \<subseteq> z \<Longrightarrow> {z} \<union> c \<in> chains S"
   416 
   446   for z :: "'a set"
   417 lemma chains_extend:
       
   418   "[| c \<in> chains S; z \<in> S; \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chains S"
       
   419   unfolding chains_def chain_subset_def by blast
   447   unfolding chains_def chain_subset_def by blast
   420 
   448 
   421 lemma mono_Chains: "r \<subseteq> s \<Longrightarrow> Chains r \<subseteq> Chains s"
   449 lemma mono_Chains: "r \<subseteq> s \<Longrightarrow> Chains r \<subseteq> Chains s"
   422   unfolding Chains_def by blast
   450   unfolding Chains_def by blast
   423 
   451 
   425   unfolding chain_subset_def subset.chain_def by fast
   453   unfolding chain_subset_def subset.chain_def by fast
   426 
   454 
   427 lemma chains_alt_def: "chains A = {C. subset.chain A C}"
   455 lemma chains_alt_def: "chains A = {C. subset.chain A C}"
   428   by (simp add: chains_def chain_subset_alt_def subset.chain_def)
   456   by (simp add: chains_def chain_subset_alt_def subset.chain_def)
   429 
   457 
   430 lemma Chains_subset:
   458 lemma Chains_subset: "Chains r \<subseteq> {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
   431   "Chains r \<subseteq> {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
       
   432   by (force simp add: Chains_def pred_on.chain_def)
   459   by (force simp add: Chains_def pred_on.chain_def)
   433 
   460 
   434 lemma Chains_subset':
   461 lemma Chains_subset':
   435   assumes "refl r"
   462   assumes "refl r"
   436   shows "{C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C} \<subseteq> Chains r"
   463   shows "{C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C} \<subseteq> Chains r"
   440 lemma Chains_alt_def:
   467 lemma Chains_alt_def:
   441   assumes "refl r"
   468   assumes "refl r"
   442   shows "Chains r = {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
   469   shows "Chains r = {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
   443   using assms Chains_subset Chains_subset' by blast
   470   using assms Chains_subset Chains_subset' by blast
   444 
   471 
   445 lemma Zorn_Lemma:
   472 lemma Zorn_Lemma: "\<forall>C\<in>chains A. \<Union>C \<in> A \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   446   "\<forall>C\<in>chains A. \<Union>C \<in> A \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
       
   447   using subset_Zorn' [of A] by (force simp: chains_alt_def)
   473   using subset_Zorn' [of A] by (force simp: chains_alt_def)
   448 
   474 
   449 lemma Zorn_Lemma2:
   475 lemma Zorn_Lemma2: "\<forall>C\<in>chains A. \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
   450   "\<forall>C\<in>chains A. \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
       
   451   using subset_Zorn [of A] by (auto simp: chains_alt_def)
   476   using subset_Zorn [of A] by (auto simp: chains_alt_def)
   452 
   477 
   453 text\<open>Various other lemmas\<close>
   478 text \<open>Various other lemmas\<close>
   454 
   479 
   455 lemma chainsD: "[| c \<in> chains S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
   480 lemma chainsD: "c \<in> chains S \<Longrightarrow> x \<in> c \<Longrightarrow> y \<in> c \<Longrightarrow> x \<subseteq> y \<or> y \<subseteq> x"
   456   unfolding chains_def chain_subset_def by blast
   481   unfolding chains_def chain_subset_def by blast
   457 
   482 
   458 lemma chainsD2: "!!(c :: 'a set set). c \<in> chains S ==> c \<subseteq> S"
   483 lemma chainsD2: "c \<in> chains S \<Longrightarrow> c \<subseteq> S"
   459   unfolding chains_def by blast
   484   unfolding chains_def by blast
   460 
   485 
   461 lemma Zorns_po_lemma:
   486 lemma Zorns_po_lemma:
   462   assumes po: "Partial_order r"
   487   assumes po: "Partial_order r"
   463     and u: "\<forall>C\<in>Chains r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a, u) \<in> r"
   488     and u: "\<forall>C\<in>Chains r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a, u) \<in> r"
   464   shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
   489   shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
   465 proof -
   490 proof -
   466   have "Preorder r" using po by (simp add: partial_order_on_def)
   491   have "Preorder r"
   467 \<comment>\<open>Mirror r in the set of subsets below (wrt r) elements of A\<close>
   492     using po by (simp add: partial_order_on_def)
   468   let ?B = "%x. r\<inverse> `` {x}" let ?S = "?B ` Field r"
   493   txt \<open>Mirror \<open>r\<close> in the set of subsets below (wrt \<open>r\<close>) elements of \<open>A\<close>.\<close>
   469   {
   494   let ?B = "\<lambda>x. r\<inverse> `` {x}"
   470     fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
   495   let ?S = "?B ` Field r"
       
   496   have "\<exists>u\<in>Field r. \<forall>A\<in>C. A \<subseteq> r\<inverse> `` {u}"  (is "\<exists>u\<in>Field r. ?P u")
       
   497     if 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A" for C
       
   498   proof -
   471     let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
   499     let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
   472     have "C = ?B ` ?A" using 1 by (auto simp: image_def)
   500     from 1 have "C = ?B ` ?A" by (auto simp: image_def)
   473     have "?A \<in> Chains r"
   501     have "?A \<in> Chains r"
   474     proof (simp add: Chains_def, intro allI impI, elim conjE)
   502     proof (simp add: Chains_def, intro allI impI, elim conjE)
   475       fix a b
   503       fix a b
   476       assume "a \<in> Field r" and "?B a \<in> C" and "b \<in> Field r" and "?B b \<in> C"
   504       assume "a \<in> Field r" and "?B a \<in> C" and "b \<in> Field r" and "?B b \<in> C"
   477       hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
   505       with 2 have "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" by auto
   478       thus "(a, b) \<in> r \<or> (b, a) \<in> r"
   506       then show "(a, b) \<in> r \<or> (b, a) \<in> r"
   479         using \<open>Preorder r\<close> and \<open>a \<in> Field r\<close> and \<open>b \<in> Field r\<close>
   507         using \<open>Preorder r\<close> and \<open>a \<in> Field r\<close> and \<open>b \<in> Field r\<close>
   480         by (simp add:subset_Image1_Image1_iff)
   508         by (simp add:subset_Image1_Image1_iff)
   481     qed
   509     qed
   482     then obtain u where uA: "u \<in> Field r" "\<forall>a\<in>?A. (a, u) \<in> r" using u by auto
   510     with u obtain u where uA: "u \<in> Field r" "\<forall>a\<in>?A. (a, u) \<in> r" by auto
   483     have "\<forall>A\<in>C. A \<subseteq> r\<inverse> `` {u}" (is "?P u")
   511     have "?P u"
   484     proof auto
   512     proof auto
   485       fix a B assume aB: "B \<in> C" "a \<in> B"
   513       fix a B assume aB: "B \<in> C" "a \<in> B"
   486       with 1 obtain x where "x \<in> Field r" and "B = r\<inverse> `` {x}" by auto
   514       with 1 obtain x where "x \<in> Field r" and "B = r\<inverse> `` {x}" by auto
   487       thus "(a, u) \<in> r" using uA and aB and \<open>Preorder r\<close>
   515       then show "(a, u) \<in> r"
       
   516         using uA and aB and \<open>Preorder r\<close>
   488         unfolding preorder_on_def refl_on_def by simp (fast dest: transD)
   517         unfolding preorder_on_def refl_on_def by simp (fast dest: transD)
   489     qed
   518     qed
   490     then have "\<exists>u\<in>Field r. ?P u" using \<open>u \<in> Field r\<close> by blast
   519     then show ?thesis
   491   }
   520       using \<open>u \<in> Field r\<close> by blast
       
   521   qed
   492   then have "\<forall>C\<in>chains ?S. \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U"
   522   then have "\<forall>C\<in>chains ?S. \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U"
   493     by (auto simp: chains_def chain_subset_def)
   523     by (auto simp: chains_def chain_subset_def)
   494   from Zorn_Lemma2 [OF this]
   524   from Zorn_Lemma2 [OF this] obtain m B
   495   obtain m B where "m \<in> Field r" and "B = r\<inverse> `` {m}"
   525     where "m \<in> Field r"
   496     and "\<forall>x\<in>Field r. B \<subseteq> r\<inverse> `` {x} \<longrightarrow> r\<inverse> `` {x} = B"
   526       and "B = r\<inverse> `` {m}"
       
   527       and "\<forall>x\<in>Field r. B \<subseteq> r\<inverse> `` {x} \<longrightarrow> r\<inverse> `` {x} = B"
   497     by auto
   528     by auto
   498   hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
   529   then have "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
   499     using po and \<open>Preorder r\<close> and \<open>m \<in> Field r\<close>
   530     using po and \<open>Preorder r\<close> and \<open>m \<in> Field r\<close>
   500     by (auto simp: subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
   531     by (auto simp: subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
   501   thus ?thesis using \<open>m \<in> Field r\<close> by blast
   532   then show ?thesis
       
   533     using \<open>m \<in> Field r\<close> by blast
   502 qed
   534 qed
   503 
   535 
   504 
   536 
   505 subsection \<open>The Well Ordering Theorem\<close>
   537 subsection \<open>The Well Ordering Theorem\<close>
   506 
   538 
   507 (* The initial segment of a relation appears generally useful.
   539 (* The initial segment of a relation appears generally useful.
   508    Move to Relation.thy?
   540    Move to Relation.thy?
   509    Definition correct/most general?
   541    Definition correct/most general?
   510    Naming?
   542    Naming?
   511 *)
   543 *)
   512 definition init_seg_of :: "(('a \<times> 'a) set \<times> ('a \<times> 'a) set) set" where
   544 definition init_seg_of :: "(('a \<times> 'a) set \<times> ('a \<times> 'a) set) set"
   513   "init_seg_of = {(r, s). r \<subseteq> s \<and> (\<forall>a b c. (a, b) \<in> s \<and> (b, c) \<in> r \<longrightarrow> (a, b) \<in> r)}"
   545   where "init_seg_of = {(r, s). r \<subseteq> s \<and> (\<forall>a b c. (a, b) \<in> s \<and> (b, c) \<in> r \<longrightarrow> (a, b) \<in> r)}"
   514 
   546 
   515 abbreviation
   547 abbreviation initial_segment_of_syntax :: "('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"
   516   initialSegmentOf :: "('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" (infix "initial'_segment'_of" 55)
   548     (infix "initial'_segment'_of" 55)
   517 where
   549   where "r initial_segment_of s \<equiv> (r, s) \<in> init_seg_of"
   518   "r initial_segment_of s \<equiv> (r, s) \<in> init_seg_of"
       
   519 
   550 
   520 lemma refl_on_init_seg_of [simp]: "r initial_segment_of r"
   551 lemma refl_on_init_seg_of [simp]: "r initial_segment_of r"
   521   by (simp add: init_seg_of_def)
   552   by (simp add: init_seg_of_def)
   522 
   553 
   523 lemma trans_init_seg_of:
   554 lemma trans_init_seg_of:
   524   "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
   555   "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
   525   by (simp (no_asm_use) add: init_seg_of_def) blast
   556   by (simp (no_asm_use) add: init_seg_of_def) blast
   526 
   557 
   527 lemma antisym_init_seg_of:
   558 lemma antisym_init_seg_of: "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r = s"
   528   "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r = s"
       
   529   unfolding init_seg_of_def by safe
   559   unfolding init_seg_of_def by safe
   530 
   560 
   531 lemma Chains_init_seg_of_Union:
   561 lemma Chains_init_seg_of_Union: "R \<in> Chains init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
   532   "R \<in> Chains init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
       
   533   by (auto simp: init_seg_of_def Ball_def Chains_def) blast
   562   by (auto simp: init_seg_of_def Ball_def Chains_def) blast
   534 
   563 
   535 lemma chain_subset_trans_Union:
   564 lemma chain_subset_trans_Union:
   536   assumes "chain\<^sub>\<subseteq> R" "\<forall>r\<in>R. trans r"
   565   assumes "chain\<^sub>\<subseteq> R" "\<forall>r\<in>R. trans r"
   537   shows "trans (\<Union>R)"
   566   shows "trans (\<Union>R)"
   538 proof (intro transI, elim UnionE)
   567 proof (intro transI, elim UnionE)
   539   fix  S1 S2 :: "'a rel" and x y z :: 'a
   568   fix S1 S2 :: "'a rel" and x y z :: 'a
   540   assume "S1 \<in> R" "S2 \<in> R"
   569   assume "S1 \<in> R" "S2 \<in> R"
   541   with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1" unfolding chain_subset_def by blast
   570   with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1"
       
   571     unfolding chain_subset_def by blast
   542   moreover assume "(x, y) \<in> S1" "(y, z) \<in> S2"
   572   moreover assume "(x, y) \<in> S1" "(y, z) \<in> S2"
   543   ultimately have "((x, y) \<in> S1 \<and> (y, z) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, z) \<in> S2)" by blast
   573   ultimately have "((x, y) \<in> S1 \<and> (y, z) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, z) \<in> S2)"
   544   with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "(x, z) \<in> \<Union>R" by (auto elim: transE)
   574     by blast
       
   575   with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "(x, z) \<in> \<Union>R"
       
   576     by (auto elim: transE)
   545 qed
   577 qed
   546 
   578 
   547 lemma chain_subset_antisym_Union:
   579 lemma chain_subset_antisym_Union:
   548   assumes "chain\<^sub>\<subseteq> R" "\<forall>r\<in>R. antisym r"
   580   assumes "chain\<^sub>\<subseteq> R" "\<forall>r\<in>R. antisym r"
   549   shows "antisym (\<Union>R)"
   581   shows "antisym (\<Union>R)"
   550 proof (intro antisymI, elim UnionE)
   582 proof (intro antisymI, elim UnionE)
   551   fix  S1 S2 :: "'a rel" and x y :: 'a
   583   fix S1 S2 :: "'a rel" and x y :: 'a
   552   assume "S1 \<in> R" "S2 \<in> R"
   584   assume "S1 \<in> R" "S2 \<in> R"
   553   with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1" unfolding chain_subset_def by blast
   585   with assms(1) have "S1 \<subseteq> S2 \<or> S2 \<subseteq> S1"
       
   586     unfolding chain_subset_def by blast
   554   moreover assume "(x, y) \<in> S1" "(y, x) \<in> S2"
   587   moreover assume "(x, y) \<in> S1" "(y, x) \<in> S2"
   555   ultimately have "((x, y) \<in> S1 \<and> (y, x) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, x) \<in> S2)" by blast
   588   ultimately have "((x, y) \<in> S1 \<and> (y, x) \<in> S1) \<or> ((x, y) \<in> S2 \<and> (y, x) \<in> S2)"
   556   with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "x = y" unfolding antisym_def by auto
   589     by blast
       
   590   with \<open>S1 \<in> R\<close> \<open>S2 \<in> R\<close> assms(2) show "x = y"
       
   591     unfolding antisym_def by auto
   557 qed
   592 qed
   558 
   593 
   559 lemma chain_subset_Total_Union:
   594 lemma chain_subset_Total_Union:
   560   assumes "chain\<^sub>\<subseteq> R" and "\<forall>r\<in>R. Total r"
   595   assumes "chain\<^sub>\<subseteq> R" and "\<forall>r\<in>R. Total r"
   561   shows "Total (\<Union>R)"
   596   shows "Total (\<Union>R)"
   562 proof (simp add: total_on_def Ball_def, auto del: disjCI)
   597 proof (simp add: total_on_def Ball_def, auto del: disjCI)
   563   fix r s a b assume A: "r \<in> R" "s \<in> R" "a \<in> Field r" "b \<in> Field s" "a \<noteq> b"
   598   fix r s a b
       
   599   assume A: "r \<in> R" "s \<in> R" "a \<in> Field r" "b \<in> Field s" "a \<noteq> b"
   564   from \<open>chain\<^sub>\<subseteq> R\<close> and \<open>r \<in> R\<close> and \<open>s \<in> R\<close> have "r \<subseteq> s \<or> s \<subseteq> r"
   600   from \<open>chain\<^sub>\<subseteq> R\<close> and \<open>r \<in> R\<close> and \<open>s \<in> R\<close> have "r \<subseteq> s \<or> s \<subseteq> r"
   565     by (auto simp add: chain_subset_def)
   601     by (auto simp add: chain_subset_def)
   566   thus "(\<exists>r\<in>R. (a, b) \<in> r) \<or> (\<exists>r\<in>R. (b, a) \<in> r)"
   602   then show "(\<exists>r\<in>R. (a, b) \<in> r) \<or> (\<exists>r\<in>R. (b, a) \<in> r)"
   567   proof
   603   proof
   568     assume "r \<subseteq> s" hence "(a, b) \<in> s \<or> (b, a) \<in> s" using assms(2) A mono_Field[of r s]
   604     assume "r \<subseteq> s"
       
   605     then have "(a, b) \<in> s \<or> (b, a) \<in> s"
       
   606       using assms(2) A mono_Field[of r s]
   569       by (auto simp add: total_on_def)
   607       by (auto simp add: total_on_def)
   570     thus ?thesis using \<open>s \<in> R\<close> by blast
   608     then show ?thesis
       
   609       using \<open>s \<in> R\<close> by blast
   571   next
   610   next
   572     assume "s \<subseteq> r" hence "(a, b) \<in> r \<or> (b, a) \<in> r" using assms(2) A mono_Field[of s r]
   611     assume "s \<subseteq> r"
       
   612     then have "(a, b) \<in> r \<or> (b, a) \<in> r"
       
   613       using assms(2) A mono_Field[of s r]
   573       by (fastforce simp add: total_on_def)
   614       by (fastforce simp add: total_on_def)
   574     thus ?thesis using \<open>r \<in> R\<close> by blast
   615     then show ?thesis
       
   616       using \<open>r \<in> R\<close> by blast
   575   qed
   617   qed
   576 qed
   618 qed
   577 
   619 
   578 lemma wf_Union_wf_init_segs:
   620 lemma wf_Union_wf_init_segs:
   579   assumes "R \<in> Chains init_seg_of" and "\<forall>r\<in>R. wf r"
   621   assumes "R \<in> Chains init_seg_of"
       
   622     and "\<forall>r\<in>R. wf r"
   580   shows "wf (\<Union>R)"
   623   shows "wf (\<Union>R)"
   581 proof(simp add: wf_iff_no_infinite_down_chain, rule ccontr, auto)
   624 proof (simp add: wf_iff_no_infinite_down_chain, rule ccontr, auto)
   582   fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f (Suc i), f i) \<in> r"
   625   fix f
       
   626   assume 1: "\<forall>i. \<exists>r\<in>R. (f (Suc i), f i) \<in> r"
   583   then obtain r where "r \<in> R" and "(f (Suc 0), f 0) \<in> r" by auto
   627   then obtain r where "r \<in> R" and "(f (Suc 0), f 0) \<in> r" by auto
   584   { fix i have "(f (Suc i), f i) \<in> r"
   628   have "(f (Suc i), f i) \<in> r" for i
   585     proof (induct i)
   629   proof (induct i)
   586       case 0 show ?case by fact
   630     case 0
   587     next
   631     show ?case by fact
   588       case (Suc i)
   632   next
   589       then obtain s where s: "s \<in> R" "(f (Suc (Suc i)), f(Suc i)) \<in> s"
   633     case (Suc i)
   590         using 1 by auto
   634     then obtain s where s: "s \<in> R" "(f (Suc (Suc i)), f(Suc i)) \<in> s"
   591       then have "s initial_segment_of r \<or> r initial_segment_of s"
   635       using 1 by auto
   592         using assms(1) \<open>r \<in> R\<close> by (simp add: Chains_def)
   636     then have "s initial_segment_of r \<or> r initial_segment_of s"
   593       with Suc s show ?case by (simp add: init_seg_of_def) blast
   637       using assms(1) \<open>r \<in> R\<close> by (simp add: Chains_def)
   594     qed
   638     with Suc s show ?case by (simp add: init_seg_of_def) blast
   595   }
   639   qed
   596   thus False using assms(2) and \<open>r \<in> R\<close>
   640   then show False
       
   641     using assms(2) and \<open>r \<in> R\<close>
   597     by (simp add: wf_iff_no_infinite_down_chain) blast
   642     by (simp add: wf_iff_no_infinite_down_chain) blast
   598 qed
   643 qed
   599 
   644 
   600 lemma initial_segment_of_Diff:
   645 lemma initial_segment_of_Diff: "p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s"
   601   "p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s"
       
   602   unfolding init_seg_of_def by blast
   646   unfolding init_seg_of_def by blast
   603 
   647 
   604 lemma Chains_inits_DiffI:
   648 lemma Chains_inits_DiffI: "R \<in> Chains init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chains init_seg_of"
   605   "R \<in> Chains init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chains init_seg_of"
       
   606   unfolding Chains_def by (blast intro: initial_segment_of_Diff)
   649   unfolding Chains_def by (blast intro: initial_segment_of_Diff)
   607 
   650 
   608 theorem well_ordering: "\<exists>r::'a rel. Well_order r \<and> Field r = UNIV"
   651 theorem well_ordering: "\<exists>r::'a rel. Well_order r \<and> Field r = UNIV"
   609 proof -
   652 proof -
   610 \<comment> \<open>The initial segment relation on well-orders:\<close>
   653 \<comment> \<open>The initial segment relation on well-orders:\<close>
   611   let ?WO = "{r::'a rel. Well_order r}"
   654   let ?WO = "{r::'a rel. Well_order r}"
   612   define I where "I = init_seg_of \<inter> ?WO \<times> ?WO"
   655   define I where "I = init_seg_of \<inter> ?WO \<times> ?WO"
   613   have I_init: "I \<subseteq> init_seg_of" by (auto simp: I_def)
   656   then have I_init: "I \<subseteq> init_seg_of" by simp
   614   hence subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
   657   then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
   615     unfolding init_seg_of_def chain_subset_def Chains_def by blast
   658     unfolding init_seg_of_def chain_subset_def Chains_def by blast
   616   have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
   659   have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
   617     by (simp add: Chains_def I_def) blast
   660     by (simp add: Chains_def I_def) blast
   618   have FI: "Field I = ?WO" by (auto simp add: I_def init_seg_of_def Field_def)
   661   have FI: "Field I = ?WO"
   619   hence 0: "Partial_order I"
   662     by (auto simp add: I_def init_seg_of_def Field_def)
       
   663   then have 0: "Partial_order I"
   620     by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
   664     by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
   621       trans_def I_def elim!: trans_init_seg_of)
   665         trans_def I_def elim!: trans_init_seg_of)
   622 \<comment> \<open>I-chains have upper bounds in ?WO wrt I: their Union\<close>
   666 \<comment> \<open>\<open>I\<close>-chains have upper bounds in \<open>?WO\<close> wrt \<open>I\<close>: their Union\<close>
   623   { fix R assume "R \<in> Chains I"
   667   have "\<Union>R \<in> ?WO \<and> (\<forall>r\<in>R. (r, \<Union>R) \<in> I)" if "R \<in> Chains I" for R
   624     hence Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
   668   proof -
   625     have subch: "chain\<^sub>\<subseteq> R" using \<open>R : Chains I\<close> I_init
   669     from that have Ris: "R \<in> Chains init_seg_of"
   626       by (auto simp: init_seg_of_def chain_subset_def Chains_def)
   670       using mono_Chains [OF I_init] by blast
       
   671     have subch: "chain\<^sub>\<subseteq> R"
       
   672       using \<open>R : Chains I\<close> I_init by (auto simp: init_seg_of_def chain_subset_def Chains_def)
   627     have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r"
   673     have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r"
   628       and "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)"
   674       and "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)"
   629       using Chains_wo [OF \<open>R \<in> Chains I\<close>] by (simp_all add: order_on_defs)
   675       using Chains_wo [OF \<open>R \<in> Chains I\<close>] by (simp_all add: order_on_defs)
   630     have "Refl (\<Union>R)" using \<open>\<forall>r\<in>R. Refl r\<close> unfolding refl_on_def by fastforce
   676     have "Refl (\<Union>R)"
       
   677       using \<open>\<forall>r\<in>R. Refl r\<close> unfolding refl_on_def by fastforce
   631     moreover have "trans (\<Union>R)"
   678     moreover have "trans (\<Union>R)"
   632       by (rule chain_subset_trans_Union [OF subch \<open>\<forall>r\<in>R. trans r\<close>])
   679       by (rule chain_subset_trans_Union [OF subch \<open>\<forall>r\<in>R. trans r\<close>])
   633     moreover have "antisym (\<Union>R)"
   680     moreover have "antisym (\<Union>R)"
   634       by (rule chain_subset_antisym_Union [OF subch \<open>\<forall>r\<in>R. antisym r\<close>])
   681       by (rule chain_subset_antisym_Union [OF subch \<open>\<forall>r\<in>R. antisym r\<close>])
   635     moreover have "Total (\<Union>R)"
   682     moreover have "Total (\<Union>R)"
   638     proof -
   685     proof -
   639       have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
   686       have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
   640       with \<open>\<forall>r\<in>R. wf (r - Id)\<close> and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
   687       with \<open>\<forall>r\<in>R. wf (r - Id)\<close> and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
   641       show ?thesis by fastforce
   688       show ?thesis by fastforce
   642     qed
   689     qed
   643     ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs)
   690     ultimately have "Well_order (\<Union>R)"
   644     moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
   691       by (simp add:order_on_defs)
   645       by(simp add: Chains_init_seg_of_Union)
   692     moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R"
   646     ultimately have "\<Union>R \<in> ?WO \<and> (\<forall>r\<in>R. (r, \<Union>R) \<in> I)"
   693       using Ris by (simp add: Chains_init_seg_of_Union)
       
   694     ultimately show ?thesis
   647       using mono_Chains [OF I_init] Chains_wo[of R] and \<open>R \<in> Chains I\<close>
   695       using mono_Chains [OF I_init] Chains_wo[of R] and \<open>R \<in> Chains I\<close>
   648       unfolding I_def by blast
   696       unfolding I_def by blast
   649   }
   697   qed
   650   hence 1: "\<forall>R \<in> Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
   698   then have 1: "\<forall>R \<in> Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I"
   651 \<comment>\<open>Zorn's Lemma yields a maximal well-order m:\<close>
   699     by (subst FI) blast
   652   then obtain m::"'a rel" where "Well_order m" and
   700 \<comment>\<open>Zorn's Lemma yields a maximal well-order \<open>m\<close>:\<close>
   653     max: "\<forall>r. Well_order r \<and> (m, r) \<in> I \<longrightarrow> r = m"
   701   then obtain m :: "'a rel"
       
   702     where "Well_order m"
       
   703       and max: "\<forall>r. Well_order r \<and> (m, r) \<in> I \<longrightarrow> r = m"
   654     using Zorns_po_lemma[OF 0 1] unfolding FI by fastforce
   704     using Zorns_po_lemma[OF 0 1] unfolding FI by fastforce
   655 \<comment>\<open>Now show by contradiction that m covers the whole type:\<close>
   705 \<comment>\<open>Now show by contradiction that \<open>m\<close> covers the whole type:\<close>
   656   { fix x::'a assume "x \<notin> Field m"
   706   have False if "x \<notin> Field m" for x :: 'a
   657 \<comment>\<open>We assume that x is not covered and extend m at the top with x\<close>
   707   proof -
       
   708 \<comment>\<open>Assuming that \<open>x\<close> is not covered and extend \<open>m\<close> at the top with \<open>x\<close>\<close>
   658     have "m \<noteq> {}"
   709     have "m \<noteq> {}"
   659     proof
   710     proof
   660       assume "m = {}"
   711       assume "m = {}"
   661       moreover have "Well_order {(x, x)}"
   712       moreover have "Well_order {(x, x)}"
   662         by (simp add: order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def)
   713         by (simp add: order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def)
   663       ultimately show False using max
   714       ultimately show False using max
   664         by (auto simp: I_def init_seg_of_def simp del: Field_insert)
   715         by (auto simp: I_def init_seg_of_def simp del: Field_insert)
   665     qed
   716     qed
   666     hence "Field m \<noteq> {}" by(auto simp:Field_def)
   717     then have "Field m \<noteq> {}" by (auto simp: Field_def)
   667     moreover have "wf (m - Id)" using \<open>Well_order m\<close>
   718     moreover have "wf (m - Id)"
   668       by (simp add: well_order_on_def)
   719       using \<open>Well_order m\<close> by (simp add: well_order_on_def)
   669 \<comment>\<open>The extension of m by x:\<close>
   720 \<comment>\<open>The extension of \<open>m\<close> by \<open>x\<close>:\<close>
   670     let ?s = "{(a, x) | a. a \<in> Field m}"
   721     let ?s = "{(a, x) | a. a \<in> Field m}"
   671     let ?m = "insert (x, x) m \<union> ?s"
   722     let ?m = "insert (x, x) m \<union> ?s"
   672     have Fm: "Field ?m = insert x (Field m)"
   723     have Fm: "Field ?m = insert x (Field m)"
   673       by (auto simp: Field_def)
   724       by (auto simp: Field_def)
   674     have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
   725     have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
   675       using \<open>Well_order m\<close> by (simp_all add: order_on_defs)
   726       using \<open>Well_order m\<close> by (simp_all add: order_on_defs)
   676 \<comment>\<open>We show that the extension is a well-order\<close>
   727 \<comment>\<open>We show that the extension is a well-order\<close>
   677     have "Refl ?m" using \<open>Refl m\<close> Fm unfolding refl_on_def by blast
   728     have "Refl ?m"
       
   729       using \<open>Refl m\<close> Fm unfolding refl_on_def by blast
   678     moreover have "trans ?m" using \<open>trans m\<close> and \<open>x \<notin> Field m\<close>
   730     moreover have "trans ?m" using \<open>trans m\<close> and \<open>x \<notin> Field m\<close>
   679       unfolding trans_def Field_def by blast
   731       unfolding trans_def Field_def by blast
   680     moreover have "antisym ?m" using \<open>antisym m\<close> and \<open>x \<notin> Field m\<close>
   732     moreover have "antisym ?m"
   681       unfolding antisym_def Field_def by blast
   733       using \<open>antisym m\<close> and \<open>x \<notin> Field m\<close> unfolding antisym_def Field_def by blast
   682     moreover have "Total ?m" using \<open>Total m\<close> and Fm by (auto simp: total_on_def)
   734     moreover have "Total ?m"
       
   735       using \<open>Total m\<close> and Fm by (auto simp: total_on_def)
   683     moreover have "wf (?m - Id)"
   736     moreover have "wf (?m - Id)"
   684     proof -
   737     proof -
   685       have "wf ?s" using \<open>x \<notin> Field m\<close> 
   738       have "wf ?s"
   686         by (auto simp: wf_eq_minimal Field_def Bex_def)
   739         using \<open>x \<notin> Field m\<close> by (auto simp: wf_eq_minimal Field_def Bex_def)
   687       thus ?thesis using \<open>wf (m - Id)\<close> and \<open>x \<notin> Field m\<close>
   740       then show ?thesis
   688         wf_subset [OF \<open>wf ?s\<close> Diff_subset]
   741         using \<open>wf (m - Id)\<close> and \<open>x \<notin> Field m\<close> wf_subset [OF \<open>wf ?s\<close> Diff_subset]
   689         by (auto simp: Un_Diff Field_def intro: wf_Un)
   742         by (auto simp: Un_Diff Field_def intro: wf_Un)
   690     qed
   743     qed
   691     ultimately have "Well_order ?m" by (simp add: order_on_defs)
   744     ultimately have "Well_order ?m"
   692 \<comment>\<open>We show that the extension is above m\<close>
   745       by (simp add: order_on_defs)
   693     moreover have "(m, ?m) \<in> I" using \<open>Well_order ?m\<close> and \<open>Well_order m\<close> and \<open>x \<notin> Field m\<close>
   746 \<comment>\<open>We show that the extension is above \<open>m\<close>\<close>
       
   747     moreover have "(m, ?m) \<in> I"
       
   748       using \<open>Well_order ?m\<close> and \<open>Well_order m\<close> and \<open>x \<notin> Field m\<close>
   694       by (fastforce simp: I_def init_seg_of_def Field_def)
   749       by (fastforce simp: I_def init_seg_of_def Field_def)
   695     ultimately
   750     ultimately
   696 \<comment>\<open>This contradicts maximality of m:\<close>
   751 \<comment>\<open>This contradicts maximality of \<open>m\<close>:\<close>
   697     have False using max and \<open>x \<notin> Field m\<close> unfolding Field_def by blast
   752     show False
   698   }
   753       using max and \<open>x \<notin> Field m\<close> unfolding Field_def by blast
   699   hence "Field m = UNIV" by auto
   754   qed
       
   755   then have "Field m = UNIV" by auto
   700   with \<open>Well_order m\<close> show ?thesis by blast
   756   with \<open>Well_order m\<close> show ?thesis by blast
   701 qed
   757 qed
   702 
   758 
   703 corollary well_order_on: "\<exists>r::'a rel. well_order_on A r"
   759 corollary well_order_on: "\<exists>r::'a rel. well_order_on A r"
   704 proof -
   760 proof -
   705   obtain r::"'a rel" where wo: "Well_order r" and univ: "Field r = UNIV"
   761   obtain r :: "'a rel" where wo: "Well_order r" and univ: "Field r = UNIV"
   706     using well_ordering [where 'a = "'a"] by blast
   762     using well_ordering [where 'a = "'a"] by blast
   707   let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
   763   let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
   708   have 1: "Field ?r = A" using wo univ
   764   have 1: "Field ?r = A"
   709     by (fastforce simp: Field_def order_on_defs refl_on_def)
   765     using wo univ by (fastforce simp: Field_def order_on_defs refl_on_def)
   710   have "Refl r" and "trans r" and "antisym r" and "Total r" and "wf (r - Id)"
   766   from \<open>Well_order r\<close> have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)"
   711     using \<open>Well_order r\<close> by (simp_all add: order_on_defs)
   767     by (simp_all add: order_on_defs)
   712   have "Refl ?r" using \<open>Refl r\<close> by (auto simp: refl_on_def 1 univ)
   768   from \<open>Refl r\<close> have "Refl ?r"
   713   moreover have "trans ?r" using \<open>trans r\<close>
   769     by (auto simp: refl_on_def 1 univ)
       
   770   moreover from \<open>trans r\<close> have "trans ?r"
   714     unfolding trans_def by blast
   771     unfolding trans_def by blast
   715   moreover have "antisym ?r" using \<open>antisym r\<close>
   772   moreover from \<open>antisym r\<close> have "antisym ?r"
   716     unfolding antisym_def by blast
   773     unfolding antisym_def by blast
   717   moreover have "Total ?r" using \<open>Total r\<close> by (simp add:total_on_def 1 univ)
   774   moreover from \<open>Total r\<close> have "Total ?r"
   718   moreover have "wf (?r - Id)" by (rule wf_subset [OF \<open>wf (r - Id)\<close>]) blast
   775     by (simp add:total_on_def 1 univ)
   719   ultimately have "Well_order ?r" by (simp add: order_on_defs)
   776   moreover have "wf (?r - Id)"
       
   777     by (rule wf_subset [OF \<open>wf (r - Id)\<close>]) blast
       
   778   ultimately have "Well_order ?r"
       
   779     by (simp add: order_on_defs)
   720   with 1 show ?thesis by auto
   780   with 1 show ?thesis by auto
   721 qed
   781 qed
   722 
   782 
   723 (* Move this to Hilbert Choice and wfrec to Wellfounded*)
   783 (* Move this to Hilbert Choice and wfrec to Wellfounded*)
   724 
   784 
   725 lemma wfrec_def_adm: "f \<equiv> wfrec R F \<Longrightarrow> wf R \<Longrightarrow> adm_wf R F \<Longrightarrow> f = F f"
   785 lemma wfrec_def_adm: "f \<equiv> wfrec R F \<Longrightarrow> wf R \<Longrightarrow> adm_wf R F \<Longrightarrow> f = F f"
   726   using wfrec_fixpoint by simp
   786   using wfrec_fixpoint by simp
   727 
   787 
   728 lemma dependent_wf_choice:
   788 lemma dependent_wf_choice:
   729   fixes P :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
   789   fixes P :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
   730   assumes "wf R" and adm: "\<And>f g x r. (\<And>z. (z, x) \<in> R \<Longrightarrow> f z = g z) \<Longrightarrow> P f x r = P g x r"
   790   assumes "wf R"
   731   assumes P: "\<And>x f. (\<And>y. (y, x) \<in> R \<Longrightarrow> P f y (f y)) \<Longrightarrow> \<exists>r. P f x r"
   791     and adm: "\<And>f g x r. (\<And>z. (z, x) \<in> R \<Longrightarrow> f z = g z) \<Longrightarrow> P f x r = P g x r"
       
   792     and P: "\<And>x f. (\<And>y. (y, x) \<in> R \<Longrightarrow> P f y (f y)) \<Longrightarrow> \<exists>r. P f x r"
   732   shows "\<exists>f. \<forall>x. P f x (f x)"
   793   shows "\<exists>f. \<forall>x. P f x (f x)"
   733 proof (intro exI allI)
   794 proof (intro exI allI)
   734   fix x 
   795   fix x
   735   define f where "f \<equiv> wfrec R (\<lambda>f x. SOME r. P f x r)"
   796   define f where "f \<equiv> wfrec R (\<lambda>f x. SOME r. P f x r)"
   736   from \<open>wf R\<close> show "P f x (f x)"
   797   from \<open>wf R\<close> show "P f x (f x)"
   737   proof (induct x)
   798   proof (induct x)
   738     fix x assume "\<And>y. (y, x) \<in> R \<Longrightarrow> P f y (f y)"
   799     case (less x)
   739     show "P f x (f x)"
   800     show "P f x (f x)"
   740     proof (subst (2) wfrec_def_adm[OF f_def \<open>wf R\<close>])
   801     proof (subst (2) wfrec_def_adm[OF f_def \<open>wf R\<close>])
   741       show "adm_wf R (\<lambda>f x. SOME r. P f x r)"
   802       show "adm_wf R (\<lambda>f x. SOME r. P f x r)"
   742         by (auto simp add: adm_wf_def intro!: arg_cong[where f=Eps] ext adm)
   803         by (auto simp add: adm_wf_def intro!: arg_cong[where f=Eps] ext adm)
   743       show "P f x (Eps (P f x))"
   804       show "P f x (Eps (P f x))"
   746   qed
   807   qed
   747 qed
   808 qed
   748 
   809 
   749 lemma (in wellorder) dependent_wellorder_choice:
   810 lemma (in wellorder) dependent_wellorder_choice:
   750   assumes "\<And>r f g x. (\<And>y. y < x \<Longrightarrow> f y = g y) \<Longrightarrow> P f x r = P g x r"
   811   assumes "\<And>r f g x. (\<And>y. y < x \<Longrightarrow> f y = g y) \<Longrightarrow> P f x r = P g x r"
   751   assumes P: "\<And>x f. (\<And>y. y < x \<Longrightarrow> P f y (f y)) \<Longrightarrow> \<exists>r. P f x r"
   812     and P: "\<And>x f. (\<And>y. y < x \<Longrightarrow> P f y (f y)) \<Longrightarrow> \<exists>r. P f x r"
   752   shows "\<exists>f. \<forall>x. P f x (f x)"
   813   shows "\<exists>f. \<forall>x. P f x (f x)"
   753   using wf by (rule dependent_wf_choice) (auto intro!: assms)
   814   using wf by (rule dependent_wf_choice) (auto intro!: assms)
   754 
   815 
   755 end
   816 end