src/HOL/Library/Zorn.thy
author wenzelm
Thu, 16 Feb 2012 22:53:24 +0100
changeset 46507 1b24c24017dd
parent 46008 c296c75f4cf4
child 46752 e9e7209eb375
permissions -rw-r--r--
tuned proofs;
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
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     2
    Author:     Jacques D. Fleuriot, Tobias Nipkow
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     3
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
     4
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
     5
The well-ordering theorem.
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
     6
*)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
     7
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
     8
header {* Zorn's Lemma *}
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    10
theory Zorn
30663
0b6aff7451b2 Main is (Complex_Main) base entry point in library theories
haftmann
parents: 30198
diff changeset
    11
imports Order_Relation Main
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14706
diff changeset
    12
begin
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    13
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
    14
(* Define globally? In Set.thy? *)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
    15
definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^bsub>\<subseteq>\<^esub>") where
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
    16
"chain\<^bsub>\<subseteq>\<^esub> C \<equiv> \<forall>A\<in>C.\<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
    17
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    18
text{*
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    19
  The lemma and section numbers refer to an unpublished article
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    20
  \cite{Abrial-Laffitte}.
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    21
*}
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    22
19736
wenzelm
parents: 18585
diff changeset
    23
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    24
  chain     ::  "'a set set => 'a set set set" where
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
    25
  "chain S  = {F. F \<subseteq> S & chain\<^bsub>\<subseteq>\<^esub> F}"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    26
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    27
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    28
  super     ::  "['a set set,'a set set] => 'a set set set" where
19736
wenzelm
parents: 18585
diff changeset
    29
  "super S c = {d. d \<in> chain S & c \<subset> d}"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    30
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    31
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    32
  maxchain  ::  "'a set set => 'a set set set" where
19736
wenzelm
parents: 18585
diff changeset
    33
  "maxchain S = {c. c \<in> chain S & super S c = {}}"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    34
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    35
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    36
  succ      ::  "['a set set,'a set set] => 'a set set" where
19736
wenzelm
parents: 18585
diff changeset
    37
  "succ S c =
wenzelm
parents: 18585
diff changeset
    38
    (if c \<notin> chain S | c \<in> maxchain S
wenzelm
parents: 18585
diff changeset
    39
    then c else SOME c'. c' \<in> super S c)"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    40
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    41
inductive_set
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    42
  TFin :: "'a set set => 'a set set set"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    43
  for S :: "'a set set"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    44
  where
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    45
    succI:        "x \<in> TFin S ==> succ S x \<in> TFin S"
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    46
  | Pow_UnionI:   "Y \<in> Pow(TFin S) ==> Union(Y) \<in> TFin S"
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    47
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    48
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    49
subsection{*Mathematical Preamble*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    50
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    51
lemma Union_lemma0:
18143
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
    52
    "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C) \<subseteq> A | B \<subseteq> Union(C)"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    53
  by blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    54
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    55
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    56
text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*}
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    57
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    58
lemma Abrial_axiom1: "x \<subseteq> succ S x"
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26298
diff changeset
    59
  apply (auto simp add: succ_def super_def maxchain_def)
18585
5d379fe2eb74 replaced swap by contrapos_np;
wenzelm
parents: 18143
diff changeset
    60
  apply (rule contrapos_np, assumption)
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    61
  apply (rule someI2)
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    62
  apply blast+
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    63
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    64
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    65
lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI]
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    66
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    67
lemma TFin_induct:
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    68
  assumes H: "n \<in> TFin S" and
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    69
    I: "!!x. x \<in> TFin S ==> P x ==> P (succ S x)"
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    70
      "!!Y. Y \<subseteq> TFin S ==> Ball Y P ==> P (Union Y)"
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    71
  shows "P n"
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
    72
  using H by induct (blast intro: I)+
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    73
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    74
lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    75
  apply (erule subset_trans)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    76
  apply (rule Abrial_axiom1)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    77
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    78
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    79
text{*Lemma 1 of section 3.1*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    80
lemma TFin_linear_lemma1:
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    81
     "[| n \<in> TFin S;  m \<in> TFin S;
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
    82
         \<forall>x \<in> TFin S. x \<subseteq> m --> x = m | succ S x \<subseteq> m
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    83
      |] ==> n \<subseteq> m | succ S m \<subseteq> n"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    84
  apply (erule TFin_induct)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    85
   apply (erule_tac [2] Union_lemma0)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    86
  apply (blast del: subsetI intro: succ_trans)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    87
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    88
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    89
text{* Lemma 2 of section 3.2 *}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    90
lemma TFin_linear_lemma2:
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
    91
     "m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    92
  apply (erule TFin_induct)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    93
   apply (rule impI [THEN ballI])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    94
   txt{*case split using @{text TFin_linear_lemma1}*}
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    95
   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    96
     assumption+)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    97
    apply (drule_tac x = n in bspec, assumption)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    98
    apply (blast del: subsetI intro: succ_trans, blast)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    99
  txt{*second induction step*}
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   100
  apply (rule impI [THEN ballI])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   101
  apply (rule Union_lemma0 [THEN disjE])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   102
    apply (rule_tac [3] disjI2)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   103
    prefer 2 apply blast
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   104
   apply (rule ballI)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   105
   apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE],
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   106
     assumption+, auto)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   107
  apply (blast intro!: Abrial_axiom1 [THEN subsetD])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   108
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   109
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   110
text{*Re-ordering the premises of Lemma 2*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   111
lemma TFin_subsetD:
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   112
     "[| n \<subseteq> m;  m \<in> TFin S;  n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   113
  by (rule TFin_linear_lemma2 [rule_format])
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   114
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   115
text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   116
lemma TFin_subset_linear: "[| m \<in> TFin S;  n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   117
  apply (rule disjE)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   118
    apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   119
      apply (assumption+, erule disjI2)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   120
  apply (blast del: subsetI
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   121
    intro: subsetI Abrial_axiom1 [THEN subset_trans])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   122
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   123
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   124
text{*Lemma 3 of section 3.3*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   125
lemma eq_succ_upper: "[| n \<in> TFin S;  m \<in> TFin S;  m = succ S m |] ==> n \<subseteq> m"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   126
  apply (erule TFin_induct)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   127
   apply (drule TFin_subsetD)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   128
     apply (assumption+, force, blast)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   129
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   130
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   131
text{*Property 3.3 of section 3.3*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   132
lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   133
  apply (rule iffI)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   134
   apply (rule Union_upper [THEN equalityI])
18143
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   135
    apply assumption
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   136
   apply (rule eq_succ_upper [THEN Union_least], assumption+)
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   137
  apply (erule ssubst)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   138
  apply (rule Abrial_axiom1 [THEN equalityI])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   139
  apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   140
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   141
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   142
subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   143
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
   144
text{*NB: We assume the partial ordering is @{text "\<subseteq>"},
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   145
 the subset relation!*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   146
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   147
lemma empty_set_mem_chain: "({} :: 'a set set) \<in> chain S"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   148
by (unfold chain_def chain_subset_def) auto
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   149
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   150
lemma super_subset_chain: "super S c \<subseteq> chain S"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   151
  by (unfold super_def) blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   152
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   153
lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   154
  by (unfold maxchain_def) blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   155
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   156
lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> EX d. d \<in> super S c"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   157
  by (unfold super_def maxchain_def) auto
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   158
18143
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   159
lemma select_super:
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   160
     "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c): super S c"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   161
  apply (erule mem_super_Ex [THEN exE])
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
   162
  apply (rule someI2)
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
   163
  apply auto
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   164
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   165
18143
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   166
lemma select_not_equals:
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   167
     "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c) \<noteq> c"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   168
  apply (rule notI)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   169
  apply (drule select_super)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26298
diff changeset
   170
  apply (simp add: super_def less_le)
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   171
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   172
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   173
lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)"
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   174
  by (unfold succ_def) (blast intro!: if_not_P)
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   175
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   176
lemma succ_not_equals: "c \<in> chain S - maxchain S ==> succ S c \<noteq> c"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   177
  apply (frule succI3)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   178
  apply (simp (no_asm_simp))
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   179
  apply (rule select_not_equals, assumption)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   180
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   181
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   182
lemma TFin_chain_lemma4: "c \<in> TFin S ==> (c :: 'a set set): chain S"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   183
  apply (erule TFin_induct)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   184
   apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]])
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   185
  apply (unfold chain_def chain_subset_def)
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   186
  apply (rule CollectI, safe)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   187
   apply (drule bspec, assumption)
46008
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
   188
   apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE])
c296c75f4cf4 reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents: 44890
diff changeset
   189
      apply blast+
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   190
  done
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
   191
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   192
theorem Hausdorff: "\<exists>c. (c :: 'a set set): maxchain S"
18143
fe14f0282c60 tidying
paulson
parents: 17200
diff changeset
   193
  apply (rule_tac x = "Union (TFin S)" in exI)
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   194
  apply (rule classical)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   195
  apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ")
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   196
   prefer 2
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   197
   apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   198
  apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4])
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   199
  apply (drule DiffI [THEN succ_not_equals], blast+)
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   200
  done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   201
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   202
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
   203
subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   204
                               There Is  a Maximal Element*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   205
14706
71590b7733b7 tuned document;
wenzelm
parents: 13652
diff changeset
   206
lemma chain_extend:
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   207
  "[| c \<in> chain S; z \<in> S; \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chain S"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   208
by (unfold chain_def chain_subset_def) blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   209
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   210
lemma chain_Union_upper: "[| c \<in> chain S; x \<in> c |] ==> x \<subseteq> Union(c)"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   211
by auto
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   212
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   213
lemma chain_ball_Union_upper: "c \<in> chain S ==> \<forall>x \<in> c. x \<subseteq> Union(c)"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   214
by auto
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   215
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   216
lemma maxchain_Zorn:
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   217
  "[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   218
apply (rule ccontr)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   219
apply (simp add: maxchain_def)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   220
apply (erule conjE)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   221
apply (subgoal_tac "({u} Un c) \<in> super S c")
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   222
 apply simp
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26298
diff changeset
   223
apply (unfold super_def less_le)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   224
apply (blast intro: chain_extend dest: chain_Union_upper)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   225
done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   226
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   227
theorem Zorn_Lemma:
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   228
  "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   229
apply (cut_tac Hausdorff maxchain_subset_chain)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   230
apply (erule exE)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   231
apply (drule subsetD, assumption)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   232
apply (drule bspec, assumption)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   233
apply (rule_tac x = "Union(c)" in bexI)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   234
 apply (rule ballI, rule impI)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   235
 apply (blast dest!: maxchain_Zorn, assumption)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   236
done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   237
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   238
subsection{*Alternative version of Zorn's Lemma*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   239
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   240
lemma Zorn_Lemma2:
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   241
  "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
   242
    ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   243
apply (cut_tac Hausdorff maxchain_subset_chain)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   244
apply (erule exE)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   245
apply (drule subsetD, assumption)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   246
apply (drule bspec, assumption, erule bexE)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   247
apply (rule_tac x = y in bexI)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   248
 prefer 2 apply assumption
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   249
apply clarify
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   250
apply (rule ccontr)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   251
apply (frule_tac z = x in chain_extend)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   252
  apply (assumption, blast)
26806
40b411ec05aa Adapted to encoding of sets as predicates
berghofe
parents: 26298
diff changeset
   253
apply (unfold maxchain_def super_def less_le)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   254
apply (blast elim!: equalityCE)
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   255
done
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   256
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   257
text{*Various other lemmas*}
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   258
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   259
lemma chainD: "[| c \<in> chain S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   260
by (unfold chain_def chain_subset_def) blast
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   261
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   262
lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   263
by (unfold chain_def) blast
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   264
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   265
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   266
(* Define globally? In Relation.thy? *)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   267
definition Chain :: "('a*'a)set \<Rightarrow> 'a set set" where
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   268
"Chain r \<equiv> {A. \<forall>a\<in>A.\<forall>b\<in>A. (a,b) : r \<or> (b,a) \<in> r}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   269
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   270
lemma mono_Chain: "r \<subseteq> s \<Longrightarrow> Chain r \<subseteq> Chain s"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   271
unfolding Chain_def by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   272
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   273
text{* Zorn's lemma for partial orders: *}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   274
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   275
lemma Zorns_po_lemma:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   276
assumes po: "Partial_order r" and u: "\<forall>C\<in>Chain r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a,u):r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   277
shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m,a):r \<longrightarrow> a=m"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   278
proof-
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   279
  have "Preorder r" using po by(simp add:partial_order_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   280
--{* Mirror r in the set of subsets below (wrt r) elements of A*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   281
  let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   282
  have "\<forall>C \<in> chain ?S. EX U:?S. ALL A:C. A\<subseteq>U"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   283
  proof (auto simp:chain_def chain_subset_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   284
    fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C.\<forall>B\<in>C. A\<subseteq>B | B\<subseteq>A"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   285
    let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   286
    have "C = ?B ` ?A" using 1 by(auto simp: image_def)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   287
    have "?A\<in>Chain r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   288
    proof (simp add:Chain_def, intro allI impI, elim conjE)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   289
      fix a b
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   290
      assume "a \<in> Field r" "?B a \<in> C" "b \<in> Field r" "?B b \<in> C"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   291
      hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   292
      thus "(a, b) \<in> r \<or> (b, a) \<in> r" using `Preorder r` `a:Field r` `b:Field r`
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   293
        by (simp add:subset_Image1_Image1_iff)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   294
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   295
    then obtain u where uA: "u:Field r" "\<forall>a\<in>?A. (a,u) : r" using u by auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   296
    have "\<forall>A\<in>C. A \<subseteq> r^-1 `` {u}" (is "?P u")
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   297
    proof auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   298
      fix a B assume aB: "B:C" "a:B"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   299
      with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   300
      thus "(a,u) : r" using uA aB `Preorder r`
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   301
        by (auto simp add: preorder_on_def refl_on_def) (metis transD)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   302
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   303
    thus "EX u:Field r. ?P u" using `u:Field r` by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   304
  qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   305
  from Zorn_Lemma2[OF this]
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   306
  obtain m B where "m:Field r" "B = r^-1 `` {m}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   307
    "\<forall>x\<in>Field r. B \<subseteq> r^-1 `` {x} \<longrightarrow> B = r^-1 `` {x}"
27064
267cab537760 Tuned proof.
ballarin
parents: 26806
diff changeset
   308
    by auto
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   309
  hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m" using po `Preorder r` `m:Field r`
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   310
    by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   311
  thus ?thesis using `m:Field r` by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   312
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   313
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   314
(* The initial segment of a relation appears generally useful.
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   315
   Move to Relation.thy?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   316
   Definition correct/most general?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   317
   Naming?
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   318
*)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   319
definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   320
"init_seg_of == {(r,s). r \<subseteq> s \<and> (\<forall>a b c. (a,b):s \<and> (b,c):r \<longrightarrow> (a,b):r)}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   321
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   322
abbreviation initialSegmentOf :: "('a*'a)set \<Rightarrow> ('a*'a)set \<Rightarrow> bool"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   323
             (infix "initial'_segment'_of" 55) where
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   324
"r initial_segment_of s == (r,s):init_seg_of"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   325
30198
922f944f03b2 name changes
nipkow
parents: 28952
diff changeset
   326
lemma refl_on_init_seg_of[simp]: "r initial_segment_of r"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   327
by(simp add:init_seg_of_def)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   328
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   329
lemma trans_init_seg_of:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   330
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   331
by(simp (no_asm_use) add: init_seg_of_def)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   332
  (metis Domain_iff UnCI Un_absorb2 subset_trans)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   333
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   334
lemma antisym_init_seg_of:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   335
  "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r=s"
35175
61255c81da01 fix more looping simp rules
huffman
parents: 32960
diff changeset
   336
unfolding init_seg_of_def by safe
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   337
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   338
lemma Chain_init_seg_of_Union:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   339
  "R \<in> Chain init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   340
by(auto simp add:init_seg_of_def Chain_def Ball_def) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   341
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   342
lemma chain_subset_trans_Union:
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   343
  "chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans(\<Union>R)"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   344
apply(auto simp add:chain_subset_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   345
apply(simp (no_asm_use) add:trans_def)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   346
apply (metis subsetD)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   347
done
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   348
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   349
lemma chain_subset_antisym_Union:
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   350
  "chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym(\<Union>R)"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   351
apply(auto simp add:chain_subset_def antisym_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   352
apply (metis subsetD)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   353
done
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   354
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   355
lemma chain_subset_Total_Union:
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   356
assumes "chain\<^bsub>\<subseteq>\<^esub> R" "\<forall>r\<in>R. Total r"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   357
shows "Total (\<Union>R)"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   358
proof (simp add: total_on_def Ball_def, auto del:disjCI)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   359
  fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a\<noteq>b"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   360
  from `chain\<^bsub>\<subseteq>\<^esub> R` `r:R` `s:R` have "r\<subseteq>s \<or> s\<subseteq>r"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   361
    by(simp add:chain_subset_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   362
  thus "(\<exists>r\<in>R. (a,b) \<in> r) \<or> (\<exists>r\<in>R. (b,a) \<in> r)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   363
  proof
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   364
    assume "r\<subseteq>s" hence "(a,b):s \<or> (b,a):s" using assms(2) A
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   365
      by(simp add:total_on_def)(metis mono_Field subsetD)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   366
    thus ?thesis using `s:R` by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   367
  next
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   368
    assume "s\<subseteq>r" hence "(a,b):r \<or> (b,a):r" using assms(2) A
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   369
      by(simp add:total_on_def)(metis mono_Field subsetD)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   370
    thus ?thesis using `r:R` by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   371
  qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   372
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   373
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   374
lemma wf_Union_wf_init_segs:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   375
assumes "R \<in> Chain init_seg_of" and "\<forall>r\<in>R. wf r" shows "wf(\<Union>R)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   376
proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   377
  fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f(Suc i), f i) \<in> r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   378
  then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   379
  { fix i have "(f(Suc i), f i) \<in> r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   380
    proof(induct i)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   381
      case 0 show ?case by fact
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   382
    next
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   383
      case (Suc i)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   384
      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
   385
        using 1 by auto
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   386
      moreover hence "s initial_segment_of r \<or> r initial_segment_of s"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   387
        using assms(1) `r:R` by(simp add: Chain_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   388
      ultimately show ?case by(simp add:init_seg_of_def) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   389
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   390
  }
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   391
  thus False using assms(2) `r:R`
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   392
    by(simp add:wf_iff_no_infinite_down_chain) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   393
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   394
27476
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   395
lemma initial_segment_of_Diff:
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   396
  "p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s"
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   397
unfolding init_seg_of_def by blast
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   398
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   399
lemma Chain_inits_DiffI:
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   400
  "R \<in> Chain init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chain init_seg_of"
27476
964766deef47 fixed extremely slow proof of Chain_inits_DiffI
huffman
parents: 27064
diff changeset
   401
unfolding Chain_def by (blast intro: initial_segment_of_Diff)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   402
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   403
theorem well_ordering: "\<exists>r::('a*'a)set. Well_order r \<and> Field r = UNIV"
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   404
proof-
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   405
-- {*The initial segment relation on well-orders: *}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   406
  let ?WO = "{r::('a*'a)set. Well_order r}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   407
  def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   408
  have I_init: "I \<subseteq> init_seg_of" by(auto simp:I_def)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   409
  hence subch: "!!R. R : Chain I \<Longrightarrow> chain\<^bsub>\<subseteq>\<^esub> R"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   410
    by(auto simp:init_seg_of_def chain_subset_def Chain_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   411
  have Chain_wo: "!!R r. R \<in> Chain I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   412
    by(simp add:Chain_def I_def) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   413
  have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   414
  hence 0: "Partial_order I"
30198
922f944f03b2 name changes
nipkow
parents: 28952
diff changeset
   415
    by(auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def trans_def I_def elim!: trans_init_seg_of)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   416
-- {*I-chains have upper bounds in ?WO wrt I: their Union*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   417
  { fix R assume "R \<in> Chain I"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   418
    hence Ris: "R \<in> Chain init_seg_of" using mono_Chain[OF I_init] by blast
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   419
    have subch: "chain\<^bsub>\<subseteq>\<^esub> R" using `R : Chain I` I_init
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   420
      by(auto simp:init_seg_of_def chain_subset_def Chain_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   421
    have "\<forall>r\<in>R. Refl r" "\<forall>r\<in>R. trans r" "\<forall>r\<in>R. antisym r" "\<forall>r\<in>R. Total r"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   422
         "\<forall>r\<in>R. wf(r-Id)"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   423
      using Chain_wo[OF `R \<in> Chain I`] by(simp_all add:order_on_defs)
30198
922f944f03b2 name changes
nipkow
parents: 28952
diff changeset
   424
    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
   425
    moreover have "trans (\<Union>R)"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   426
      by(rule chain_subset_trans_Union[OF subch `\<forall>r\<in>R. trans r`])
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   427
    moreover have "antisym(\<Union>R)"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   428
      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
   429
    moreover have "Total (\<Union>R)"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   430
      by(rule chain_subset_Total_Union[OF subch `\<forall>r\<in>R. Total r`])
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   431
    moreover have "wf((\<Union>R)-Id)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   432
    proof-
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   433
      have "(\<Union>R)-Id = \<Union>{r-Id|r. r \<in> R}" by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   434
      with `\<forall>r\<in>R. wf(r-Id)` wf_Union_wf_init_segs[OF Chain_inits_DiffI[OF Ris]]
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   435
      show ?thesis by (simp (no_asm_simp)) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   436
    qed
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   437
    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
   438
    moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   439
      by(simp add: Chain_init_seg_of_Union)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   440
    ultimately have "\<Union>R : ?WO \<and> (\<forall>r\<in>R. (r,\<Union>R) : I)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   441
      using mono_Chain[OF I_init] `R \<in> Chain I`
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   442
      by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo subsetD)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   443
  }
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   444
  hence 1: "\<forall>R \<in> Chain I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r,u) : I" by (subst FI) blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   445
--{*Zorn's Lemma yields a maximal well-order m:*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   446
  then obtain m::"('a*'a)set" where "Well_order m" and
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   447
    max: "\<forall>r. Well_order r \<and> (m,r):I \<longrightarrow> r=m"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   448
    using Zorns_po_lemma[OF 0 1] by (auto simp:FI)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   449
--{*Now show by contradiction that m covers the whole type:*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   450
  { fix x::'a assume "x \<notin> Field m"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   451
--{*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
   452
    have "m \<noteq> {}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   453
    proof
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   454
      assume "m={}"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   455
      moreover have "Well_order {(x,x)}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   456
        by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_def Range_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   457
      ultimately show False using max
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   458
        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
   459
    qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   460
    hence "Field m \<noteq> {}" by(auto simp:Field_def)
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   461
    moreover have "wf(m-Id)" using `Well_order m`
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   462
      by(simp add:well_order_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   463
--{*The extension of m by x:*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   464
    let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   465
    have Fm: "Field ?m = insert x (Field m)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   466
      apply(simp add:Field_insert Field_Un)
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   467
      unfolding Field_def by auto
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   468
    have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   469
      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
   470
--{*We show that the extension is a well-order*}
30198
922f944f03b2 name changes
nipkow
parents: 28952
diff changeset
   471
    have "Refl ?m" using `Refl m` Fm by(auto simp:refl_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   472
    moreover have "trans ?m" using `trans m` `x \<notin> Field m`
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   473
      unfolding trans_def Field_def Domain_def Range_def by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   474
    moreover have "antisym ?m" using `antisym m` `x \<notin> Field m`
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   475
      unfolding antisym_def Field_def Domain_def Range_def by blast
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   476
    moreover have "Total ?m" using `Total m` Fm by(auto simp: total_on_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   477
    moreover have "wf(?m-Id)"
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   478
    proof-
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   479
      have "wf ?s" using `x \<notin> Field m`
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   480
        by(auto simp add:wf_eq_minimal Field_def Domain_def Range_def) metis
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   481
      thus ?thesis using `wf(m-Id)` `x \<notin> Field m`
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30663
diff changeset
   482
        wf_subset[OF `wf ?s` Diff_subset]
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 35175
diff changeset
   483
        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
   484
    qed
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   485
    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
   486
--{*We show that the extension is above m*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   487
    moreover hence "(m,?m) : I" using `Well_order m` `x \<notin> Field m`
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 35175
diff changeset
   488
      by(fastforce simp:I_def init_seg_of_def Field_def Domain_def Range_def)
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   489
    ultimately
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   490
--{*This contradicts maximality of m:*}
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   491
    have False using max `x \<notin> Field m` unfolding Field_def by blast
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   492
  }
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   493
  hence "Field m = UNIV" by auto
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   494
  moreover with `Well_order m` have "Well_order m" by simp
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   495
  ultimately show ?thesis by blast
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   496
qed
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   497
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   498
corollary well_order_on: "\<exists>r::('a*'a)set. well_order_on A r"
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   499
proof -
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   500
  obtain r::"('a*'a)set" where wo: "Well_order r" and univ: "Field r = UNIV"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   501
    using well_ordering[where 'a = "'a"] by blast
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   502
  let ?r = "{(x,y). x:A & y:A & (x,y):r}"
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   503
  have 1: "Field ?r = A" using wo univ
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 35175
diff changeset
   504
    by(fastforce simp: Field_def Domain_def Range_def order_on_defs refl_on_def)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   505
  have "Refl r" "trans r" "antisym r" "Total r" "wf(r-Id)"
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   506
    using `Well_order r` by(simp_all add:order_on_defs)
30198
922f944f03b2 name changes
nipkow
parents: 28952
diff changeset
   507
  have "Refl ?r" using `Refl r` by(auto simp:refl_on_def 1 univ)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   508
  moreover have "trans ?r" using `trans r`
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   509
    unfolding trans_def by blast
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   510
  moreover have "antisym ?r" using `antisym r`
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   511
    unfolding antisym_def by blast
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   512
  moreover have "Total ?r" using `Total r` by(simp add:total_on_def 1 univ)
26272
d63776c3be97 Added Order_Relation
nipkow
parents: 26191
diff changeset
   513
  moreover have "wf(?r - Id)" by(rule wf_subset[OF `wf(r-Id)`]) blast
26295
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   514
  ultimately have "Well_order ?r" by(simp add:order_on_defs)
afc43168ed85 More defns and thms
nipkow
parents: 26272
diff changeset
   515
  with 1 show ?thesis by metis
26191
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   516
qed
ae537f315b34 Generalized Zorn and added well-ordering theorem
nipkow
parents: 25691
diff changeset
   517
13551
b7f64ee8da84 converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff changeset
   518
end