src/HOL/Library/Zorn.thy
author Christian Sternagel
Fri, 09 Aug 2013 19:34:23 +0900
changeset 53007 54e290da6da8
parent 52821 05eb2d77b195
child 53374 a14d2a854c02
permissions -rw-r--r--
avoid low-level Same structure;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     1
(*  Title:      HOL/Library/Zorn.thy
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
     2
    Author:     Jacques D. Fleuriot
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
     3
    Author:     Tobias Nipkow, TUM
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
     4
    Author:     Christian Sternagel, JAIST
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     5
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     6
Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     7
The well-ordering theorem.
52199
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
     8
The extension of any well-founded relation to a well-order. 
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
     9
*)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    10
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    11
header {* Zorn's Lemma *}
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    12
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    13
theory Zorn
52199
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
    14
imports Order_Union
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    15
begin
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    16
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    17
subsection {* Zorn's Lemma for the Subset Relation *}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    18
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    19
subsubsection {* Results that do not require an order *}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    20
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    21
text {*Let @{text P} be a binary predicate on the set @{text A}.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    22
locale pred_on =
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    23
  fixes A :: "'a set"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    24
    and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubset>" 50)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    25
begin
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    26
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    27
abbreviation Peq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<sqsubseteq>" 50) where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    28
  "x \<sqsubseteq> y \<equiv> P\<^sup>=\<^sup>= x y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    29
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    30
text {*A chain is a totally ordered subset of @{term A}.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    31
definition chain :: "'a set \<Rightarrow> bool" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    32
  "chain C \<longleftrightarrow> C \<subseteq> A \<and> (\<forall>x\<in>C. \<forall>y\<in>C. x \<sqsubseteq> y \<or> y \<sqsubseteq> x)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    33
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    34
text {*We call a chain that is a proper superset of some set @{term X},
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    35
but not necessarily a chain itself, a superchain of @{term X}.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    36
abbreviation superchain :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infix "<c" 50) where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    37
  "X <c C \<equiv> chain C \<and> X \<subset> C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    38
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    39
text {*A maximal chain is a chain that does not have a superchain.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    40
definition maxchain :: "'a set \<Rightarrow> bool" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    41
  "maxchain C \<longleftrightarrow> chain C \<and> \<not> (\<exists>S. C <c S)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    42
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    43
text {*We define the successor of a set to be an arbitrary
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    44
superchain, if such exists, or the set itself, otherwise.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    45
definition suc :: "'a set \<Rightarrow> 'a set" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    46
  "suc C = (if \<not> chain C \<or> maxchain C then C else (SOME D. C <c D))"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    47
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    48
lemma chainI [Pure.intro?]:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    49
  "\<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"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    50
  unfolding chain_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    51
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    52
lemma chain_total:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    53
  "chain C \<Longrightarrow> x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    54
  by (simp add: chain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    55
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    56
lemma not_chain_suc [simp]: "\<not> chain X \<Longrightarrow> suc X = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    57
  by (simp add: suc_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    58
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    59
lemma maxchain_suc [simp]: "maxchain X \<Longrightarrow> suc X = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    60
  by (simp add: suc_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    61
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    62
lemma suc_subset: "X \<subseteq> suc X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    63
  by (auto simp: suc_def maxchain_def intro: someI2)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    64
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    65
lemma chain_empty [simp]: "chain {}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    66
  by (auto simp: chain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    67
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    68
lemma not_maxchain_Some:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    69
  "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> C <c (SOME D. C <c D)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    70
  by (rule someI_ex) (auto simp: maxchain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    71
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    72
lemma suc_not_equals:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    73
  "chain C \<Longrightarrow> \<not> maxchain C \<Longrightarrow> suc C \<noteq> C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    74
  by (auto simp: suc_def) (metis less_irrefl not_maxchain_Some)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    75
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    76
lemma subset_suc:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    77
  assumes "X \<subseteq> Y" shows "X \<subseteq> suc Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    78
  using assms by (rule subset_trans) (rule suc_subset)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    79
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    80
text {*We build a set @{term \<C>} that is closed under applications
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    81
of @{term suc} and contains the union of all its subsets.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    82
inductive_set suc_Union_closed ("\<C>") where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    83
  suc: "X \<in> \<C> \<Longrightarrow> suc X \<in> \<C>" |
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    84
  Union [unfolded Pow_iff]: "X \<in> Pow \<C> \<Longrightarrow> \<Union>X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    85
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    86
text {*Since the empty set as well as the set itself is a subset of
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    87
every set, @{term \<C>} contains at least @{term "{} \<in> \<C>"} and
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    88
@{term "\<Union>\<C> \<in> \<C>"}.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    89
lemma
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    90
  suc_Union_closed_empty: "{} \<in> \<C>" and
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    91
  suc_Union_closed_Union: "\<Union>\<C> \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    92
  using Union [of "{}"] and Union [of "\<C>"] by simp+
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    93
text {*Thus closure under @{term suc} will hit a maximal chain
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    94
eventually, as is shown below.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    95
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    96
lemma suc_Union_closed_induct [consumes 1, case_names suc Union,
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    97
  induct pred: suc_Union_closed]:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    98
  assumes "X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
    99
    and "\<And>X. \<lbrakk>X \<in> \<C>; Q X\<rbrakk> \<Longrightarrow> Q (suc X)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   100
    and "\<And>X. \<lbrakk>X \<subseteq> \<C>; \<forall>x\<in>X. Q x\<rbrakk> \<Longrightarrow> Q (\<Union>X)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   101
  shows "Q X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   102
  using assms by (induct) blast+
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   103
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   104
lemma suc_Union_closed_cases [consumes 1, case_names suc Union,
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   105
  cases pred: suc_Union_closed]:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   106
  assumes "X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   107
    and "\<And>Y. \<lbrakk>X = suc Y; Y \<in> \<C>\<rbrakk> \<Longrightarrow> Q"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   108
    and "\<And>Y. \<lbrakk>X = \<Union>Y; Y \<subseteq> \<C>\<rbrakk> \<Longrightarrow> Q"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   109
  shows "Q"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   110
  using assms by (cases) simp+
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   111
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   112
text {*On chains, @{term suc} yields a chain.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   113
lemma chain_suc:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   114
  assumes "chain X" shows "chain (suc X)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   115
  using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   116
  by (cases "\<not> chain X \<or> maxchain X")
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   117
     (force simp: suc_def dest: not_maxchain_Some)+
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   118
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   119
lemma chain_sucD:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   120
  assumes "chain X" shows "suc X \<subseteq> A \<and> chain (suc X)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   121
proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   122
  from `chain X` have "chain (suc X)" by (rule chain_suc)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   123
  moreover then have "suc X \<subseteq> A" unfolding chain_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   124
  ultimately show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   125
qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   126
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   127
lemma suc_Union_closed_total':
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   128
  assumes "X \<in> \<C>" and "Y \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   129
    and *: "\<And>Z. Z \<in> \<C> \<Longrightarrow> Z \<subseteq> Y \<Longrightarrow> Z = Y \<or> suc Z \<subseteq> Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   130
  shows "X \<subseteq> Y \<or> suc Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   131
  using `X \<in> \<C>`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   132
proof (induct)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   133
  case (suc X)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   134
  with * show ?case by (blast del: subsetI intro: subset_suc)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   135
qed blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   136
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   137
lemma suc_Union_closed_subsetD:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   138
  assumes "Y \<subseteq> X" and "X \<in> \<C>" and "Y \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   139
  shows "X = Y \<or> suc Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   140
  using assms(2-, 1)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   141
proof (induct arbitrary: Y)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   142
  case (suc X)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   143
  note * = `\<And>Y. \<lbrakk>Y \<in> \<C>; Y \<subseteq> X\<rbrakk> \<Longrightarrow> X = Y \<or> suc Y \<subseteq> X`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   144
  with suc_Union_closed_total' [OF `Y \<in> \<C>` `X \<in> \<C>`]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   145
    have "Y \<subseteq> X \<or> suc X \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   146
  then show ?case
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   147
  proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   148
    assume "Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   149
    with * and `Y \<in> \<C>` have "X = Y \<or> suc Y \<subseteq> X" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   150
    then show ?thesis
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   151
    proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   152
      assume "X = Y" then show ?thesis by simp
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   153
    next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   154
      assume "suc Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   155
      then have "suc Y \<subseteq> suc X" by (rule subset_suc)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   156
      then show ?thesis by simp
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   157
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   158
  next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   159
    assume "suc X \<subseteq> Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   160
    with `Y \<subseteq> suc X` show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   161
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   162
next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   163
  case (Union X)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   164
  show ?case
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   165
  proof (rule ccontr)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   166
    assume "\<not> ?thesis"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   167
    with `Y \<subseteq> \<Union>X` obtain x y z
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   168
    where "\<not> suc Y \<subseteq> \<Union>X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   169
      and "x \<in> X" and "y \<in> x" and "y \<notin> Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   170
      and "z \<in> suc Y" and "\<forall>x\<in>X. z \<notin> x" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   171
    with `X \<subseteq> \<C>` have "x \<in> \<C>" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   172
    from Union and `x \<in> X`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   173
      have *: "\<And>y. \<lbrakk>y \<in> \<C>; y \<subseteq> x\<rbrakk> \<Longrightarrow> x = y \<or> suc y \<subseteq> x" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   174
    with suc_Union_closed_total' [OF `Y \<in> \<C>` `x \<in> \<C>`]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   175
      have "Y \<subseteq> x \<or> suc x \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   176
    then show False
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   177
    proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   178
      assume "Y \<subseteq> x"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   179
      with * [OF `Y \<in> \<C>`] have "x = Y \<or> suc Y \<subseteq> x" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   180
      then show False
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   181
      proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   182
        assume "x = Y" with `y \<in> x` and `y \<notin> Y` show False by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   183
      next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   184
        assume "suc Y \<subseteq> x"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   185
        with `x \<in> X` have "suc Y \<subseteq> \<Union>X" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   186
        with `\<not> suc Y \<subseteq> \<Union>X` show False by contradiction
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   187
      qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   188
    next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   189
      assume "suc x \<subseteq> Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   190
      moreover from suc_subset and `y \<in> x` have "y \<in> suc x" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   191
      ultimately show False using `y \<notin> Y` by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   192
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   193
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   194
qed
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   195
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   196
text {*The elements of @{term \<C>} are totally ordered by the subset relation.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   197
lemma suc_Union_closed_total:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   198
  assumes "X \<in> \<C>" and "Y \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   199
  shows "X \<subseteq> Y \<or> Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   200
proof (cases "\<forall>Z\<in>\<C>. Z \<subseteq> Y \<longrightarrow> Z = Y \<or> suc Z \<subseteq> Y")
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   201
  case True
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   202
  with suc_Union_closed_total' [OF assms]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   203
    have "X \<subseteq> Y \<or> suc Y \<subseteq> X" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   204
  then show ?thesis using suc_subset [of Y] by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   205
next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   206
  case False
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   207
  then obtain Z
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   208
    where "Z \<in> \<C>" and "Z \<subseteq> Y" and "Z \<noteq> Y" and "\<not> suc Z \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   209
  with suc_Union_closed_subsetD and `Y \<in> \<C>` show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   210
qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   211
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   212
text {*Once we hit a fixed point w.r.t. @{term suc}, all other elements
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   213
of @{term \<C>} are subsets of this fixed point.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   214
lemma suc_Union_closed_suc:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   215
  assumes "X \<in> \<C>" and "Y \<in> \<C>" and "suc Y = Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   216
  shows "X \<subseteq> Y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   217
using `X \<in> \<C>`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   218
proof (induct)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   219
  case (suc X)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   220
  with `Y \<in> \<C>` and suc_Union_closed_subsetD
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   221
    have "X = Y \<or> suc X \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   222
  then show ?case by (auto simp: `suc Y = Y`)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   223
qed blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   224
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   225
lemma eq_suc_Union:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   226
  assumes "X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   227
  shows "suc X = X \<longleftrightarrow> X = \<Union>\<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   228
proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   229
  assume "suc X = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   230
  with suc_Union_closed_suc [OF suc_Union_closed_Union `X \<in> \<C>`]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   231
    have "\<Union>\<C> \<subseteq> X" .
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   232
  with `X \<in> \<C>` show "X = \<Union>\<C>" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   233
next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   234
  from `X \<in> \<C>` have "suc X \<in> \<C>" by (rule suc)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   235
  then have "suc X \<subseteq> \<Union>\<C>" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   236
  moreover assume "X = \<Union>\<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   237
  ultimately have "suc X \<subseteq> X" by simp
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   238
  moreover have "X \<subseteq> suc X" by (rule suc_subset)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   239
  ultimately show "suc X = X" ..
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   240
qed
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   241
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   242
lemma suc_in_carrier:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   243
  assumes "X \<subseteq> A"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   244
  shows "suc X \<subseteq> A"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   245
  using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   246
  by (cases "\<not> chain X \<or> maxchain X")
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   247
     (auto dest: chain_sucD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   248
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   249
lemma suc_Union_closed_in_carrier:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   250
  assumes "X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   251
  shows "X \<subseteq> A"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   252
  using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   253
  by (induct) (auto dest: suc_in_carrier)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   254
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   255
text {*All elements of @{term \<C>} are chains.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   256
lemma suc_Union_closed_chain:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   257
  assumes "X \<in> \<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   258
  shows "chain X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   259
using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   260
proof (induct)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   261
  case (suc X) then show ?case by (simp add: suc_def) (metis not_maxchain_Some)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   262
next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   263
  case (Union X)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   264
  then have "\<Union>X \<subseteq> A" by (auto dest: suc_Union_closed_in_carrier)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   265
  moreover have "\<forall>x\<in>\<Union>X. \<forall>y\<in>\<Union>X. x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   266
  proof (intro ballI)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   267
    fix x y
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   268
    assume "x \<in> \<Union>X" and "y \<in> \<Union>X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   269
    then obtain u v where "x \<in> u" and "u \<in> X" and "y \<in> v" and "v \<in> X" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   270
    with Union have "u \<in> \<C>" and "v \<in> \<C>" and "chain u" and "chain v" by blast+
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   271
    with suc_Union_closed_total have "u \<subseteq> v \<or> v \<subseteq> u" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   272
    then show "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   273
    proof
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   274
      assume "u \<subseteq> v"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   275
      from `chain v` show ?thesis
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   276
      proof (rule chain_total)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   277
        show "y \<in> v" by fact
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   278
        show "x \<in> v" using `u \<subseteq> v` and `x \<in> u` by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   279
      qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   280
    next
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   281
      assume "v \<subseteq> u"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   282
      from `chain u` show ?thesis
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   283
      proof (rule chain_total)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   284
        show "x \<in> u" by fact
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   285
        show "y \<in> u" using `v \<subseteq> u` and `y \<in> v` by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   286
      qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   287
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   288
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   289
  ultimately show ?case unfolding chain_def ..
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   290
qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   291
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   292
subsubsection {* Hausdorff's Maximum Principle *}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   293
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   294
text {*There exists a maximal totally ordered subset of @{term A}. (Note that we do not
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   295
require @{term A} to be partially ordered.)*}
46980
6bc213e90401 tuned specifications
haftmann
parents: 46752
diff changeset
   296
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   297
theorem Hausdorff: "\<exists>C. maxchain C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   298
proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   299
  let ?M = "\<Union>\<C>"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   300
  have "maxchain ?M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   301
  proof (rule ccontr)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   302
    assume "\<not> maxchain ?M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   303
    then have "suc ?M \<noteq> ?M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   304
      using suc_not_equals and
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   305
      suc_Union_closed_chain [OF suc_Union_closed_Union] by simp
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   306
    moreover have "suc ?M = ?M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   307
      using eq_suc_Union [OF suc_Union_closed_Union] by simp
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   308
    ultimately show False by contradiction
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   309
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   310
  then show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   311
qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   312
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   313
text {*Make notation @{term \<C>} available again.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   314
no_notation suc_Union_closed ("\<C>")
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   315
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   316
lemma chain_extend:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   317
  "chain C \<Longrightarrow> z \<in> A \<Longrightarrow> \<forall>x\<in>C. x \<sqsubseteq> z \<Longrightarrow> chain ({z} \<union> C)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   318
  unfolding chain_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   319
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   320
lemma maxchain_imp_chain:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   321
  "maxchain C \<Longrightarrow> chain C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   322
  by (simp add: maxchain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   323
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   324
end
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   325
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   326
text {*Hide constant @{const pred_on.suc_Union_closed}, which was just needed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   327
for the proof of Hausforff's maximum principle.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   328
hide_const pred_on.suc_Union_closed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   329
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   330
lemma chain_mono:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   331
  assumes "\<And>x y. \<lbrakk>x \<in> A; y \<in> A; P x y\<rbrakk> \<Longrightarrow> Q x y"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   332
    and "pred_on.chain A P C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   333
  shows "pred_on.chain A Q C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   334
  using assms unfolding pred_on.chain_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   335
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   336
subsubsection {* Results for the proper subset relation *}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   337
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   338
interpretation subset: pred_on "A" "op \<subset>" for A .
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   339
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   340
lemma subset_maxchain_max:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   341
  assumes "subset.maxchain A C" and "X \<in> A" and "\<Union>C \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   342
  shows "\<Union>C = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   343
proof (rule ccontr)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   344
  let ?C = "{X} \<union> C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   345
  from `subset.maxchain A C` have "subset.chain A C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   346
    and *: "\<And>S. subset.chain A S \<Longrightarrow> \<not> C \<subset> S"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   347
    by (auto simp: subset.maxchain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   348
  moreover have "\<forall>x\<in>C. x \<subseteq> X" using `\<Union>C \<subseteq> X` by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   349
  ultimately have "subset.chain A ?C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   350
    using subset.chain_extend [of A C X] and `X \<in> A` by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   351
  moreover assume "\<Union>C \<noteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   352
  moreover then have "C \<subset> ?C" using `\<Union>C \<subseteq> X` by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   353
  ultimately show False using * by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   354
qed
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   355
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   356
subsubsection {* Zorn's lemma *}
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   357
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   358
text {*If every chain has an upper bound, then there is a maximal set.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   359
lemma subset_Zorn:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   360
  assumes "\<And>C. subset.chain A C \<Longrightarrow> \<exists>U\<in>A. \<forall>X\<in>C. X \<subseteq> U"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   361
  shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   362
proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   363
  from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   364
  then have "subset.chain A M" by (rule subset.maxchain_imp_chain)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   365
  with assms obtain Y where "Y \<in> A" and "\<forall>X\<in>M. X \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   366
  moreover have "\<forall>X\<in>A. Y \<subseteq> X \<longrightarrow> Y = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   367
  proof (intro ballI impI)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   368
    fix X
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   369
    assume "X \<in> A" and "Y \<subseteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   370
    show "Y = X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   371
    proof (rule ccontr)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   372
      assume "Y \<noteq> X"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   373
      with `Y \<subseteq> X` have "\<not> X \<subseteq> Y" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   374
      from subset.chain_extend [OF `subset.chain A M` `X \<in> A`] and `\<forall>X\<in>M. X \<subseteq> Y`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   375
        have "subset.chain A ({X} \<union> M)" using `Y \<subseteq> X` by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   376
      moreover have "M \<subset> {X} \<union> M" using `\<forall>X\<in>M. X \<subseteq> Y` and `\<not> X \<subseteq> Y` by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   377
      ultimately show False
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   378
        using `subset.maxchain A M` by (auto simp: subset.maxchain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   379
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   380
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   381
  ultimately show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   382
qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   383
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   384
text{*Alternative version of Zorn's lemma for the subset relation.*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   385
lemma subset_Zorn':
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   386
  assumes "\<And>C. subset.chain A C \<Longrightarrow> \<Union>C \<in> A"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   387
  shows "\<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   388
proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   389
  from subset.Hausdorff [of A] obtain M where "subset.maxchain A M" ..
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   390
  then have "subset.chain A M" by (rule subset.maxchain_imp_chain)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   391
  with assms have "\<Union>M \<in> A" .
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   392
  moreover have "\<forall>Z\<in>A. \<Union>M \<subseteq> Z \<longrightarrow> \<Union>M = Z"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   393
  proof (intro ballI impI)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   394
    fix Z
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   395
    assume "Z \<in> A" and "\<Union>M \<subseteq> Z"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   396
    with subset_maxchain_max [OF `subset.maxchain A M`]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   397
      show "\<Union>M = Z" .
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   398
  qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   399
  ultimately show ?thesis by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   400
qed
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   401
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   402
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   403
subsection {* Zorn's Lemma for Partial Orders *}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   404
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   405
text {*Relate old to new definitions.*}
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   406
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   407
(* Define globally? In Set.thy? *)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   408
definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^sub>\<subseteq>") where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   409
  "chain\<^sub>\<subseteq> C \<longleftrightarrow> (\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A)"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   410
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   411
definition chains :: "'a set set \<Rightarrow> 'a set set set" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   412
  "chains A = {C. C \<subseteq> A \<and> chain\<^sub>\<subseteq> C}"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   413
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   414
(* Define globally? In Relation.thy? *)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   415
definition Chains :: "('a \<times> 'a) set \<Rightarrow> 'a set set" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   416
  "Chains r = {C. \<forall>a\<in>C. \<forall>b\<in>C. (a, b) \<in> r \<or> (b, a) \<in> r}"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   417
52183
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   418
lemma chains_extend:
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   419
  "[| c \<in> chains S; z \<in> S; \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chains S"
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   420
  by (unfold chains_def chain_subset_def) blast
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   421
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   422
lemma mono_Chains: "r \<subseteq> s \<Longrightarrow> Chains r \<subseteq> Chains s"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   423
  unfolding Chains_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   424
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   425
lemma chain_subset_alt_def: "chain\<^sub>\<subseteq> C = subset.chain UNIV C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   426
  by (auto simp add: chain_subset_def subset.chain_def)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   427
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   428
lemma chains_alt_def: "chains A = {C. subset.chain A C}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   429
  by (simp add: chains_def chain_subset_alt_def subset.chain_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   430
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   431
lemma Chains_subset:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   432
  "Chains r \<subseteq> {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   433
  by (force simp add: Chains_def pred_on.chain_def)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   434
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   435
lemma Chains_subset':
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   436
  assumes "refl r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   437
  shows "{C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C} \<subseteq> Chains r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   438
  using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   439
  by (auto simp add: Chains_def pred_on.chain_def refl_on_def)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   440
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   441
lemma Chains_alt_def:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   442
  assumes "refl r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   443
  shows "Chains r = {C. pred_on.chain UNIV (\<lambda>x y. (x, y) \<in> r) C}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   444
  using assms
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   445
  by (metis Chains_subset Chains_subset' subset_antisym)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   446
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   447
lemma Zorn_Lemma:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   448
  "\<forall>C\<in>chains A. \<Union>C \<in> A \<Longrightarrow> \<exists>M\<in>A. \<forall>X\<in>A. M \<subseteq> X \<longrightarrow> X = M"
52183
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   449
  using subset_Zorn' [of A] by (force simp: chains_alt_def)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   450
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   451
lemma Zorn_Lemma2:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   452
  "\<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"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   453
  using subset_Zorn [of A] by (auto simp: chains_alt_def)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   454
52183
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   455
text{*Various other lemmas*}
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   456
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   457
lemma chainsD: "[| c \<in> chains S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   458
by (unfold chains_def chain_subset_def) blast
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   459
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   460
lemma chainsD2: "!!(c :: 'a set set). c \<in> chains S ==> c \<subseteq> S"
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   461
by (unfold chains_def) blast
667961fa6a60 fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents: 52181
diff changeset
   462
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   463
lemma Zorns_po_lemma:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   464
  assumes po: "Partial_order r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   465
    and u: "\<forall>C\<in>Chains r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a, u) \<in> r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   466
  shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   467
proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   468
  have "Preorder r" using po by (simp add: partial_order_on_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   469
--{* Mirror r in the set of subsets below (wrt r) elements of A*}
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   470
  let ?B = "%x. r\<inverse> `` {x}" let ?S = "?B ` Field r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   471
  {
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   472
    fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C. \<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   473
    let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   474
    have "C = ?B ` ?A" using 1 by (auto simp: image_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   475
    have "?A \<in> Chains r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   476
    proof (simp add: Chains_def, intro allI impI, elim conjE)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   477
      fix a b
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   478
      assume "a \<in> Field r" and "?B a \<in> C" and "b \<in> Field r" and "?B b \<in> C"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   479
      hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   480
      thus "(a, b) \<in> r \<or> (b, a) \<in> r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   481
        using `Preorder r` and `a \<in> Field r` and `b \<in> Field r`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   482
        by (simp add:subset_Image1_Image1_iff)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   483
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   484
    then obtain u where uA: "u \<in> Field r" "\<forall>a\<in>?A. (a, u) \<in> r" using u by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   485
    have "\<forall>A\<in>C. A \<subseteq> r\<inverse> `` {u}" (is "?P u")
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   486
    proof auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   487
      fix a B assume aB: "B \<in> C" "a \<in> B"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   488
      with 1 obtain x where "x \<in> Field r" and "B = r\<inverse> `` {x}" by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   489
      thus "(a, u) \<in> r" using uA and aB and `Preorder r`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   490
        by (auto simp add: preorder_on_def refl_on_def) (metis transD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   491
    qed
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   492
    then have "\<exists>u\<in>Field r. ?P u" using `u \<in> Field r` by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   493
  }
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   494
  then have "\<forall>C\<in>chains ?S. \<exists>U\<in>?S. \<forall>A\<in>C. A \<subseteq> U"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   495
    by (auto simp: chains_def chain_subset_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   496
  from Zorn_Lemma2 [OF this]
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   497
  obtain m B where "m \<in> Field r" and "B = r\<inverse> `` {m}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   498
    and "\<forall>x\<in>Field r. B \<subseteq> r\<inverse> `` {x} \<longrightarrow> r\<inverse> `` {x} = B"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   499
    by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   500
  hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   501
    using po and `Preorder r` and `m \<in> Field r`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   502
    by (auto simp: subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   503
  thus ?thesis using `m \<in> Field r` by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   504
qed
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   505
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   506
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   507
subsection {* The Well Ordering Theorem *}
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   508
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   509
(* The initial segment of a relation appears generally useful.
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   510
   Move to Relation.thy?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   511
   Definition correct/most general?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   512
   Naming?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   513
*)
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   514
definition init_seg_of :: "(('a \<times> 'a) set \<times> ('a \<times> 'a) set) set" where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   515
  "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)}"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   516
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   517
abbreviation
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   518
  initialSegmentOf :: "('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" (infix "initial'_segment'_of" 55)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   519
where
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   520
  "r initial_segment_of s \<equiv> (r, s) \<in> init_seg_of"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   521
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   522
lemma refl_on_init_seg_of [simp]: "r initial_segment_of r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   523
  by (simp add: init_seg_of_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   524
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   525
lemma trans_init_seg_of:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   526
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   527
  by (simp (no_asm_use) add: init_seg_of_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   528
     (metis UnCI Un_absorb2 subset_trans)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   529
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   530
lemma antisym_init_seg_of:
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   531
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r = s"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   532
  unfolding init_seg_of_def by safe
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   533
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   534
lemma Chains_init_seg_of_Union:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   535
  "R \<in> Chains init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   536
  by (auto simp: init_seg_of_def Ball_def Chains_def) blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   537
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   538
lemma chain_subset_trans_Union:
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   539
  "chain\<^sub>\<subseteq> R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans (\<Union>R)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   540
apply (auto simp add: chain_subset_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   541
apply (simp (no_asm_use) add: trans_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   542
apply (metis subsetD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   543
done
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   544
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   545
lemma chain_subset_antisym_Union:
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   546
  "chain\<^sub>\<subseteq> R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym (\<Union>R)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   547
apply (auto simp add: chain_subset_def antisym_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   548
apply (metis subsetD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   549
done
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   550
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   551
lemma chain_subset_Total_Union:
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   552
  assumes "chain\<^sub>\<subseteq> R" and "\<forall>r\<in>R. Total r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   553
  shows "Total (\<Union>R)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   554
proof (simp add: total_on_def Ball_def, auto del: disjCI)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   555
  fix r s a b assume A: "r \<in> R" "s \<in> R" "a \<in> Field r" "b \<in> Field s" "a \<noteq> b"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   556
  from `chain\<^sub>\<subseteq> R` and `r \<in> R` and `s \<in> R` have "r \<subseteq> s \<or> s \<subseteq> r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   557
    by (auto simp add: chain_subset_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   558
  thus "(\<exists>r\<in>R. (a, b) \<in> r) \<or> (\<exists>r\<in>R. (b, a) \<in> r)"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   559
  proof
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   560
    assume "r \<subseteq> s" hence "(a, b) \<in> s \<or> (b, a) \<in> s" using assms(2) A
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   561
      by (simp add: total_on_def) (metis mono_Field subsetD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   562
    thus ?thesis using `s \<in> R` by blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   563
  next
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   564
    assume "s \<subseteq> r" hence "(a, b) \<in> r \<or> (b, a) \<in> r" using assms(2) A
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   565
      by (simp add: total_on_def) (metis mono_Field subsetD)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   566
    thus ?thesis using `r \<in> R` by blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   567
  qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   568
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   569
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   570
lemma wf_Union_wf_init_segs:
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   571
  assumes "R \<in> Chains init_seg_of" and "\<forall>r\<in>R. wf r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   572
  shows "wf (\<Union>R)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   573
proof(simp add: wf_iff_no_infinite_down_chain, rule ccontr, auto)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   574
  fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f (Suc i), f i) \<in> r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   575
  then obtain r where "r \<in> R" and "(f (Suc 0), f 0) \<in> r" by auto
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   576
  { fix i have "(f (Suc i), f i) \<in> r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   577
    proof (induct i)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   578
      case 0 show ?case by fact
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   579
    next
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   580
      case (Suc i)
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   581
      moreover obtain s where "s \<in> R" and "(f (Suc (Suc i)), f(Suc i)) \<in> s"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   582
        using 1 by auto
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   583
      moreover hence "s initial_segment_of r \<or> r initial_segment_of s"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   584
        using assms(1) `r \<in> R` by (simp add: Chains_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   585
      ultimately show ?case by (simp add: init_seg_of_def) blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   586
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   587
  }
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   588
  thus False using assms(2) and `r \<in> R`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   589
    by (simp add: wf_iff_no_infinite_down_chain) blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   590
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   591
27476
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   592
lemma initial_segment_of_Diff:
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   593
  "p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   594
  unfolding init_seg_of_def by blast
27476
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   595
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   596
lemma Chains_inits_DiffI:
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   597
  "R \<in> Chains init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chains init_seg_of"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   598
  unfolding Chains_def by (blast intro: initial_segment_of_Diff)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   599
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   600
theorem well_ordering: "\<exists>r::'a rel. Well_order r \<and> Field r = UNIV"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   601
proof -
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   602
-- {*The initial segment relation on well-orders: *}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   603
  let ?WO = "{r::'a rel. Well_order r}"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   604
  def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   605
  have I_init: "I \<subseteq> init_seg_of" by (auto simp: I_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   606
  hence subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   607
    by (auto simp: init_seg_of_def chain_subset_def Chains_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   608
  have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   609
    by (simp add: Chains_def I_def) blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   610
  have FI: "Field I = ?WO" by (auto simp add: I_def init_seg_of_def Field_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   611
  hence 0: "Partial_order I"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   612
    by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   613
      trans_def I_def elim!: trans_init_seg_of)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   614
-- {*I-chains have upper bounds in ?WO wrt I: their Union*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   615
  { fix R assume "R \<in> Chains I"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   616
    hence Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   617
    have subch: "chain\<^sub>\<subseteq> R" using `R : Chains I` I_init
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   618
      by (auto simp: init_seg_of_def chain_subset_def Chains_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   619
    have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   620
      and "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   621
      using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   622
    have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by (auto simp: refl_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   623
    moreover have "trans (\<Union>R)"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   624
      by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   625
    moreover have "antisym (\<Union>R)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   626
      by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   627
    moreover have "Total (\<Union>R)"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   628
      by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   629
    moreover have "wf ((\<Union>R) - Id)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   630
    proof -
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   631
      have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   632
      with `\<forall>r\<in>R. wf (r - Id)` and wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   633
      show ?thesis by (simp (no_asm_simp)) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   634
    qed
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   635
    ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   636
    moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   637
      by(simp add: Chains_init_seg_of_Union)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   638
    ultimately have "\<Union>R \<in> ?WO \<and> (\<forall>r\<in>R. (r, \<Union>R) \<in> I)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   639
      using mono_Chains [OF I_init] and `R \<in> Chains I`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   640
      by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   641
  }
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   642
  hence 1: "\<forall>R \<in> Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   643
--{*Zorn's Lemma yields a maximal well-order m:*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   644
  then obtain m::"'a rel" where "Well_order m" and
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   645
    max: "\<forall>r. Well_order r \<and> (m, r) \<in> I \<longrightarrow> r = m"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   646
    using Zorns_po_lemma[OF 0 1] by (auto simp:FI)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   647
--{*Now show by contradiction that m covers the whole type:*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   648
  { fix x::'a assume "x \<notin> Field m"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   649
--{*We assume that x is not covered and extend m at the top with x*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   650
    have "m \<noteq> {}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   651
    proof
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   652
      assume "m = {}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   653
      moreover have "Well_order {(x, x)}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   654
        by (simp add: order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   655
      ultimately show False using max
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   656
        by (auto simp: I_def init_seg_of_def simp del: Field_insert)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   657
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   658
    hence "Field m \<noteq> {}" by(auto simp:Field_def)
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   659
    moreover have "wf (m - Id)" using `Well_order m`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   660
      by (simp add: well_order_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   661
--{*The extension of m by x:*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   662
    let ?s = "{(a, x) | a. a \<in> Field m}"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   663
    let ?m = "insert (x, x) m \<union> ?s"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   664
    have Fm: "Field ?m = insert x (Field m)"
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   665
      by (auto simp: Field_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   666
    have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   667
      using `Well_order m` by (simp_all add: order_on_defs)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   668
--{*We show that the extension is a well-order*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   669
    have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   670
    moreover have "trans ?m" using `trans m` and `x \<notin> Field m`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   671
      unfolding trans_def Field_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   672
    moreover have "antisym ?m" using `antisym m` and `x \<notin> Field m`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   673
      unfolding antisym_def Field_def by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   674
    moreover have "Total ?m" using `Total m` and Fm by (auto simp: total_on_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   675
    moreover have "wf (?m - Id)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   676
    proof -
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   677
      have "wf ?s" using `x \<notin> Field m`
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   678
        by (auto simp add: wf_eq_minimal Field_def) metis
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   679
      thus ?thesis using `wf (m - Id)` and `x \<notin> Field m`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   680
        wf_subset [OF `wf ?s` Diff_subset]
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 35175
diff changeset
   681
        by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   682
    qed
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   683
    ultimately have "Well_order ?m" by (simp add: order_on_defs)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   684
--{*We show that the extension is above m*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   685
    moreover hence "(m, ?m) \<in> I" using `Well_order m` and `x \<notin> Field m`
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   686
      by (fastforce simp: I_def init_seg_of_def Field_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   687
    ultimately
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   688
--{*This contradicts maximality of m:*}
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   689
    have False using max and `x \<notin> Field m` unfolding Field_def by blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   690
  }
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   691
  hence "Field m = UNIV" by auto
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   692
  moreover with `Well_order m` have "Well_order m" by simp
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   693
  ultimately show ?thesis by blast
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   694
qed
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   695
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   696
corollary well_order_on: "\<exists>r::'a rel. well_order_on A r"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   697
proof -
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   698
  obtain r::"'a rel" where wo: "Well_order r" and univ: "Field r = UNIV"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   699
    using well_ordering [where 'a = "'a"] by blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   700
  let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   701
  have 1: "Field ?r = A" using wo univ
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   702
    by (fastforce simp: Field_def order_on_defs refl_on_def)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   703
  have "Refl r" and "trans r" and "antisym r" and "Total r" and "wf (r - Id)"
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   704
    using `Well_order r` by (simp_all add: order_on_defs)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   705
  have "Refl ?r" using `Refl r` by (auto simp: refl_on_def 1 univ)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   706
  moreover have "trans ?r" using `trans r`
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   707
    unfolding trans_def by blast
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   708
  moreover have "antisym ?r" using `antisym r`
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   709
    unfolding antisym_def by blast
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   710
  moreover have "Total ?r" using `Total r` by (simp add:total_on_def 1 univ)
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   711
  moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf (r - Id)`]) blast
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   712
  ultimately have "Well_order ?r" by (simp add: order_on_defs)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   713
  with 1 show ?thesis by metis
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   714
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   715
52199
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   716
subsection {* Extending Well-founded Relations to Well-Orders *}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   717
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   718
text {*A \emph{downset} (also lower set, decreasing set, initial segment, or
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   719
downward closed set) is closed w.r.t.\ smaller elements.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   720
definition downset_on where
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   721
  "downset_on A r = (\<forall>x y. (x, y) \<in> r \<and> y \<in> A \<longrightarrow> x \<in> A)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   722
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   723
(*
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   724
text {*Connection to order filters of the @{theory Cardinals} theory.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   725
lemma (in wo_rel) ofilter_downset_on_conv:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   726
  "ofilter A \<longleftrightarrow> downset_on A r \<and> A \<subseteq> Field r"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   727
  by (auto simp: downset_on_def ofilter_def under_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   728
*)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   729
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   730
lemma downset_onI:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   731
  "(\<And>x y. (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A) \<Longrightarrow> downset_on A r"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   732
  by (auto simp: downset_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   733
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   734
lemma downset_onD:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   735
  "downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   736
  by (auto simp: downset_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   737
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   738
text {*Extensions of relations w.r.t.\ a given set.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   739
definition extension_on where
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   740
  "extension_on A r s = (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> s \<longrightarrow> (x, y) \<in> r)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   741
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   742
lemma extension_onI:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   743
  "(\<And>x y. \<lbrakk>x \<in> A; y \<in> A; (x, y) \<in> s\<rbrakk> \<Longrightarrow> (x, y) \<in> r) \<Longrightarrow> extension_on A r s"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   744
  by (auto simp: extension_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   745
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   746
lemma extension_onD:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   747
  "extension_on A r s \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> (x, y) \<in> s \<Longrightarrow> (x, y) \<in> r"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   748
  by (auto simp: extension_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   749
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   750
lemma downset_on_Union:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   751
  assumes "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   752
  shows "downset_on (Field (\<Union>R)) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   753
  using assms by (auto intro: downset_onI dest: downset_onD)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   754
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   755
lemma chain_subset_extension_on_Union:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   756
  assumes "chain\<^sub>\<subseteq> R" and "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   757
  shows "extension_on (Field (\<Union>R)) (\<Union>R) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   758
  using assms
52821
05eb2d77b195 tuned proof;
wenzelm
parents: 52199
diff changeset
   759
  by (simp add: chain_subset_def extension_on_def) (metis mono_Field set_mp)
52199
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   760
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   761
lemma downset_on_empty [simp]: "downset_on {} p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   762
  by (auto simp: downset_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   763
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   764
lemma extension_on_empty [simp]: "extension_on {} p q"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   765
  by (auto simp: extension_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   766
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   767
text {*Every well-founded relation can be extended to a well-order.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   768
theorem well_order_extension:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   769
  assumes "wf p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   770
  shows "\<exists>w. p \<subseteq> w \<and> Well_order w"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   771
proof -
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   772
  let ?K = "{r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p}"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   773
  def I \<equiv> "init_seg_of \<inter> ?K \<times> ?K"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   774
  have I_init: "I \<subseteq> init_seg_of" by (simp add: I_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   775
  then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   776
    by (auto simp: init_seg_of_def chain_subset_def Chains_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   777
  have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow>
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   778
      Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   779
    by (simp add: Chains_def I_def) blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   780
  have FI: "Field I = ?K" by (auto simp: I_def init_seg_of_def Field_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   781
  then have 0: "Partial_order I"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   782
    by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   783
      trans_def I_def elim: trans_init_seg_of)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   784
  { fix R assume "R \<in> Chains I"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   785
    then have Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   786
    have subch: "chain\<^sub>\<subseteq> R" using `R \<in> Chains I` I_init
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   787
      by (auto simp: init_seg_of_def chain_subset_def Chains_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   788
    have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r" and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   789
      "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)" and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   790
      "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   791
      "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   792
      using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   793
    have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by (auto simp: refl_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   794
    moreover have "trans (\<Union>R)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   795
      by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   796
    moreover have "antisym (\<Union>R)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   797
      by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   798
    moreover have "Total (\<Union>R)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   799
      by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   800
    moreover have "wf ((\<Union>R) - Id)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   801
    proof -
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   802
      have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   803
      with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   804
      show ?thesis by (simp (no_asm_simp)) blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   805
    qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   806
    ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   807
    moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   808
      by (simp add: Chains_init_seg_of_Union)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   809
    moreover have "downset_on (Field (\<Union>R)) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   810
      by (rule downset_on_Union [OF `\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p`])
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   811
    moreover have "extension_on (Field (\<Union>R)) (\<Union>R) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   812
      by (rule chain_subset_extension_on_Union [OF subch `\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p`])
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   813
    ultimately have "\<Union>R \<in> ?K \<and> (\<forall>r\<in>R. (r,\<Union>R) \<in> I)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   814
      using mono_Chains [OF I_init] and `R \<in> Chains I`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   815
      by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   816
  }
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   817
  then have 1: "\<forall>R\<in>Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   818
  txt {*Zorn's Lemma yields a maximal well-order m.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   819
  from Zorns_po_lemma [OF 0 1] obtain m :: "('a \<times> 'a) set"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   820
    where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   821
    max: "\<forall>r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p \<and>
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   822
      (m, r) \<in> I \<longrightarrow> r = m"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   823
    by (auto simp: FI)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   824
  have "Field p \<subseteq> Field m"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   825
  proof (rule ccontr)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   826
    let ?Q = "Field p - Field m"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   827
    assume "\<not> (Field p \<subseteq> Field m)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   828
    with assms [unfolded wf_eq_minimal, THEN spec, of ?Q]
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   829
      obtain x where "x \<in> Field p" and "x \<notin> Field m" and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   830
      min: "\<forall>y. (y, x) \<in> p \<longrightarrow> y \<notin> ?Q" by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   831
    txt {*Add @{term x} as topmost element to @{term m}.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   832
    let ?s = "{(y, x) | y. y \<in> Field m}"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   833
    let ?m = "insert (x, x) m \<union> ?s"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   834
    have Fm: "Field ?m = insert x (Field m)" by (auto simp: Field_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   835
    have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   836
      using `Well_order m` by (simp_all add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   837
    txt {*We show that the extension is a well-order.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   838
    have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   839
    moreover have "trans ?m" using `trans m` `x \<notin> Field m`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   840
      unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   841
    moreover have "antisym ?m" using `antisym m` `x \<notin> Field m`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   842
      unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   843
    moreover have "Total ?m" using `Total m` Fm by (auto simp: Relation.total_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   844
    moreover have "wf (?m - Id)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   845
    proof -
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   846
      have "wf ?s" using `x \<notin> Field m`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   847
        by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   848
      thus ?thesis using `wf (m - Id)` `x \<notin> Field m`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   849
        wf_subset [OF `wf ?s` Diff_subset]
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   850
        by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   851
    qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   852
    ultimately have "Well_order ?m" by (simp add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   853
    moreover have "extension_on (Field ?m) ?m p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   854
      using `extension_on (Field m) m p` `downset_on (Field m) p`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   855
      by (subst Fm) (auto simp: extension_on_def dest: downset_onD)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   856
    moreover have "downset_on (Field ?m) p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   857
      using `downset_on (Field m) p` and min
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   858
      by (subst Fm, simp add: downset_on_def Field_def) (metis Domain_iff)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   859
    moreover have "(m, ?m) \<in> I"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   860
      using `Well_order m` and `Well_order ?m` and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   861
      `downset_on (Field m) p` and `downset_on (Field ?m) p` and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   862
      `extension_on (Field m) m p` and `extension_on (Field ?m) ?m p` and
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   863
      `Refl m` and `x \<notin> Field m`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   864
      by (auto simp: I_def init_seg_of_def refl_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   865
    ultimately
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   866
    --{*This contradicts maximality of m:*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   867
    show False using max and `x \<notin> Field m` unfolding Field_def by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   868
  qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   869
  have "p \<subseteq> m"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   870
    using `Field p \<subseteq> Field m` and `extension_on (Field m) m p`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   871
    by (force simp: Field_def extension_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   872
  with `Well_order m` show ?thesis by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   873
qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   874
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   875
text {*Every well-founded relation can be extended to a total well-order.*}
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   876
corollary total_well_order_extension:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   877
  assumes "wf p"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   878
  shows "\<exists>w. p \<subseteq> w \<and> Well_order w \<and> Field w = UNIV"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   879
proof -
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   880
  from well_order_extension [OF assms] obtain w
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   881
    where "p \<subseteq> w" and wo: "Well_order w" by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   882
  let ?A = "UNIV - Field w"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   883
  from well_order_on [of ?A] obtain w' where wo': "well_order_on ?A w'" ..
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   884
  have [simp]: "Field w' = ?A" using rel.well_order_on_Well_order [OF wo'] by simp
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   885
  have *: "Field w \<inter> Field w' = {}" by simp
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   886
  let ?w = "w \<union>o w'"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   887
  have "p \<subseteq> ?w" using `p \<subseteq> w` by (auto simp: Osum_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   888
  moreover have "Well_order ?w" using Osum_Well_order [OF * wo] and wo' by simp
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   889
  moreover have "Field ?w = UNIV" by (simp add: Field_Osum)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   890
  ultimately show ?thesis by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   891
qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   892
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   893
corollary well_order_on_extension:
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   894
  assumes "wf p" and "Field p \<subseteq> A"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   895
  shows "\<exists>w. p \<subseteq> w \<and> well_order_on A w"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   896
proof -
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   897
  from total_well_order_extension [OF `wf p`] obtain r
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   898
    where "p \<subseteq> r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   899
  let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   900
  from `p \<subseteq> r` have "p \<subseteq> ?r" using `Field p \<subseteq> A` by (auto simp: Field_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   901
  have 1: "Field ?r = A" using wo univ
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   902
    by (fastforce simp: Field_def order_on_defs refl_on_def)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   903
  have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)"
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   904
    using `Well_order r` by (simp_all add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   905
  have "refl_on A ?r" using `Refl r` by (auto simp: refl_on_def univ)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   906
  moreover have "trans ?r" using `trans r`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   907
    unfolding trans_def by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   908
  moreover have "antisym ?r" using `antisym r`
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   909
    unfolding antisym_def by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   910
  moreover have "total_on A ?r" using `Total r` by (simp add: total_on_def univ)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   911
  moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf(r - Id)`]) blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   912
  ultimately have "well_order_on A ?r" by (simp add: order_on_defs)
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   913
  with `p \<subseteq> ?r` show ?thesis by blast
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   914
qed
d25fc4c0ff62 merged Well_Order_Extension into Zorn
popescua
parents: 52183
diff changeset
   915
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   916
end
52181
59e5dd7b8f9a modernized Zorn (by Christian Sternagel)
popescua
parents: 51500
diff changeset
   917