src/HOL/Real/HahnBanach/ZornLemma.thy
author wenzelm
Sat, 16 Dec 2000 21:41:51 +0100
changeset 10687 c186279eecea
parent 9035 371f023d3dbd
child 13515 a6a7025fd7e8
permissions -rw-r--r--
tuned HOL/Real/HahnBanach;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/ZornLemma.thy
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     4
*)
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     5
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
     6
header {* Zorn's Lemma *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     7
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
     8
theory ZornLemma = Aux + Zorn:
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     9
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    10
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    11
  Zorn's Lemmas states: if every linear ordered subset of an ordered
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    12
  set @{text S} has an upper bound in @{text S}, then there exists a
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    13
  maximal element in @{text S}.  In our application, @{text S} is a
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    14
  set of sets ordered by set inclusion. Since the union of a chain of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    15
  sets is an upper bound for all elements of the chain, the conditions
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    16
  of Zorn's lemma can be modified: if @{text S} is non-empty, it
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    17
  suffices to show that for every non-empty chain @{text c} in @{text
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    18
  S} the union of @{text c} also lies in @{text S}.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    19
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    20
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    21
theorem Zorn's_Lemma:
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    22
  "(\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S) \<Longrightarrow> a \<in> S
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    23
  \<Longrightarrow> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    24
proof (rule Zorn_Lemma2)
8104
wenzelm
parents: 7978
diff changeset
    25
  txt_raw {* \footnote{See
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    26
  \url{http://isabelle.in.tum.de/library/HOL/HOL-Real/Zorn.html}} \isanewline *}
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    27
  assume r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    28
  assume aS: "a \<in> S"
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    29
  show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    30
  proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    31
    fix c assume "c \<in> chain S"
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    32
    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    33
    proof cases
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    34
 
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    35
      txt {* If @{text c} is an empty chain, then every element in
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    36
      @{text S} is an upper bound of @{text c}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    37
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    38
      assume "c = {}" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    39
      with aS show ?thesis by fast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    40
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    41
      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    42
      bound of @{text c}, lying in @{text S}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    43
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    44
    next
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    45
      assume c: "c \<noteq> {}"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    46
      show ?thesis 
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    47
      proof 
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    48
        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    49
        show "\<Union>c \<in> S" 
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    50
        proof (rule r)
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    51
          from c show "\<exists>x. x \<in> c" by fast  
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    52
        qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    53
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    54
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    55
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    56
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    57
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    58
end