src/HOL/HahnBanach/Bounds.thy
author haftmann
Mon, 29 Dec 2008 14:08:08 +0100
changeset 29197 6d4cb27ed19c
parent 29043 src/HOL/Real/HahnBanach/Bounds.thy@409d1ca929b5
child 29252 ea97aa6aeba2
permissions -rw-r--r--
adapted HOL source structure to distribution layout
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     1
(*  Title:      HOL/Real/HahnBanach/Bounds.thy
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    ID:         $Id$
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     4
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     5
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8659
diff changeset
     6
header {* Bounds *}
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     7
25596
ad9e3594f3f3 tuned header
haftmann
parents: 21404
diff changeset
     8
theory Bounds
29043
409d1ca929b5 fixed import: requires ContNotDenum;
wenzelm
parents: 27612
diff changeset
     9
imports Main ContNotDenum
25596
ad9e3594f3f3 tuned header
haftmann
parents: 21404
diff changeset
    10
begin
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    11
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    12
locale lub =
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    13
  fixes A and x
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    14
  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    15
    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    16
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    17
lemmas [elim?] = lub.least lub.upper
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    18
19736
wenzelm
parents: 16417
diff changeset
    19
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21210
diff changeset
    20
  the_lub :: "'a::order set \<Rightarrow> 'a" where
19736
wenzelm
parents: 16417
diff changeset
    21
  "the_lub A = The (lub A)"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 13515
diff changeset
    22
21210
c17fd2df4e9e renamed 'const_syntax' to 'notation';
wenzelm
parents: 19736
diff changeset
    23
notation (xsymbols)
19736
wenzelm
parents: 16417
diff changeset
    24
  the_lub  ("\<Squnion>_" [90] 90)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    25
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    26
lemma the_lub_equality [elim?]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    27
  assumes "lub A x"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    28
  shows "\<Squnion>A = (x::'a::order)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    29
proof -
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    30
  interpret lub [A x] by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    31
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    32
  proof (unfold the_lub_def)
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    33
    from `lub A x` show "The (lub A) = x"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    34
    proof
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    35
      fix x' assume lub': "lub A x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    36
      show "x' = x"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    37
      proof (rule order_antisym)
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    38
	from lub' show "x' \<le> x"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    39
	proof
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    40
          fix a assume "a \<in> A"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    41
          then show "a \<le> x" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    42
	qed
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    43
	show "x \<le> x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    44
	proof
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    45
          fix a assume "a \<in> A"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    46
          with lub' show "a \<le> x'" ..
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    47
	qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    48
      qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    49
    qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    50
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    51
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    52
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    53
lemma the_lubI_ex:
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    54
  assumes ex: "\<exists>x. lub A x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    55
  shows "lub A (\<Squnion>A)"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    56
proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    57
  from ex obtain x where x: "lub A x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    58
  also from x have [symmetric]: "\<Squnion>A = x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    59
  finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    60
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    61
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    62
lemma lub_compat: "lub A x = isLub UNIV A x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    63
proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    64
  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    65
    by (rule ext) (simp only: isUb_def)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    66
  then show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    67
    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8659
diff changeset
    68
qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    69
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    70
lemma real_complete:
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    71
  fixes A :: "real set"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    72
  assumes nonempty: "\<exists>a. a \<in> A"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    73
    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    74
  shows "\<exists>x. lub A x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    75
proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    76
  from ex_upper have "\<exists>y. isUb UNIV A y"
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    77
    unfolding isUb_def setle_def by blast
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    78
  with nonempty have "\<exists>x. isLub UNIV A x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    79
    by (rule reals_complete)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    80
  then show ?thesis by (simp only: lub_compat)
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    81
qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    82
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9503
diff changeset
    83
end