src/HOL/HahnBanach/ZornLemma.thy
author wenzelm
Mon, 16 Mar 2009 18:24:30 +0100
changeset 30549 d2d7874648bd
parent 29252 ea97aa6aeba2
permissions -rw-r--r--
simplified method setup;
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
    Author:     Gertrud Bauer, TU Munich
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     3
*)
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     4
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
     5
header {* Zorn's Lemma *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     6
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
     7
theory ZornLemma
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
     8
imports Zorn
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
     9
begin
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    10
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    11
text {*
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    12
  Zorn's Lemmas states: if every linear ordered subset of an ordered
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    13
  set @{text S} has an upper bound in @{text S}, then there exists a
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    14
  maximal element in @{text S}.  In our application, @{text S} is a
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    15
  set of sets ordered by set inclusion. Since the union of a chain of
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    16
  sets is an upper bound for all elements of the chain, the conditions
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    17
  of Zorn's lemma can be modified: if @{text S} is non-empty, it
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    18
  suffices to show that for every non-empty chain @{text c} in @{text
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    19
  S} the union of @{text c} also lies in @{text S}.
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    20
*}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    21
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    22
theorem Zorn's_Lemma:
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    23
  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
    24
    and aS: "a \<in> S"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    25
  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
    26
proof (rule Zorn_Lemma2)
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    27
  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
    28
  proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    29
    fix c assume "c \<in> chain S"
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    30
    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    31
    proof cases
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    32
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    33
      txt {* If @{text c} is an empty chain, then every element in
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
    34
	@{text S} is an upper bound of @{text c}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    35
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    36
      assume "c = {}"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    37
      with aS show ?thesis by fast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    38
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    39
      txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
    40
	bound of @{text c}, lying in @{text S}. *}
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    41
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    42
    next
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
    43
      assume "c \<noteq> {}"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    44
      show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    45
      proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    46
        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
    47
        show "\<Union>c \<in> S"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    48
        proof (rule r)
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
    49
          from `c \<noteq> {}` show "\<exists>x. x \<in> c" by fast
23378
1d138d6bb461 tuned proofs: avoid implicit prems;
wenzelm
parents: 16417
diff changeset
    50
	  show "c \<in> chain S" by fact
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