src/HOL/Real/HahnBanach/ZornLemma.thy
author bauerg
Fri, 07 May 2004 12:16:57 +0200
changeset 14710 247615bfffb8
parent 14024 213dcc39358f
child 16417 9bc16273c2d4
permissions -rw-r--r--
replaced Aux.thy by RealLemmas.thy
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
14710
247615bfffb8 replaced Aux.thy by RealLemmas.thy
bauerg
parents: 14024
diff changeset
     8
theory ZornLemma = 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:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    22
  assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    23
    and aS: "a \<in> S"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    24
  shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    25
proof (rule Zorn_Lemma2)
8104
wenzelm
parents: 7978
diff changeset
    26
  txt_raw {* \footnote{See
14024
213dcc39358f HOL-Real -> HOL-Complex
kleing
parents: 13515
diff changeset
    27
  \url{http://isabelle.in.tum.de/library/HOL/HOL-Complex/Zorn.html}} \isanewline *}
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    28
  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
    29
  proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    30
    fix c assume "c \<in> chain S"
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    31
    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    32
    proof cases
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    33
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    34
      txt {* If @{text c} is an empty chain, then every element in
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    35
      @{text S} is an upper bound of @{text c}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    36
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    37
      assume "c = {}"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    38
      with aS show ?thesis by fast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    39
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    40
      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    41
      bound of @{text c}, lying in @{text S}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    42
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    43
    next
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    44
      assume c: "c \<noteq> {}"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    45
      show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    46
      proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    47
        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    48
        show "\<Union>c \<in> S"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    49
        proof (rule r)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    50
          from c show "\<exists>x. x \<in> c" by fast
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    51
        qed
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
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    56
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    57
end