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