src/HOL/Real/HahnBanach/Bounds.thy
author wenzelm
Fri, 22 Oct 1999 20:14:31 +0200
changeset 7917 5e5b9813cce7
parent 7808 fd019ac3485f
child 7978 1b99ee57d131
permissions -rw-r--r--
HahnBanach update by Gertrud Bauer;
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
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
header {* Bounds *};
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     7
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     8
theory Bounds = Main + Real:;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     9
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    10
text_raw {* \begin{comment} *};
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    11
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    12
subsection {* The sets of lower and upper bounds *};
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    13
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    14
constdefs
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    15
  is_LowerBound :: "('a::order) set => 'a set => 'a => bool"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    16
  "is_LowerBound A B == \<lambda>x. x:A & (ALL y:B. x <= y)"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    17
   
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    18
  LowerBounds :: "('a::order) set => 'a set => 'a set"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    19
  "LowerBounds A B == Collect (is_LowerBound A B)"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    20
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    21
  is_UpperBound :: "('a::order) set => 'a set => 'a => bool"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    22
  "is_UpperBound A B == \<lambda>x. x:A & (ALL y:B. y <= x)"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    23
 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    24
  UpperBounds :: "('a::order) set => 'a set => 'a set"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    25
  "UpperBounds A B == Collect (is_UpperBound A B)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    26
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    27
syntax
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    28
  "_UPPERS" :: "[pttrn, 'a set, 'a => bool] => 'a set"     
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    29
    ("(3UPPER'_BOUNDS _:_./ _)" 10)
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    30
  "_UPPERS_U" :: "[pttrn, 'a => bool] => 'a set"           
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    31
    ("(3UPPER'_BOUNDS _./ _)" 10)
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    32
  "_LOWERS" :: "[pttrn, 'a set, 'a => bool] => 'a set"     
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    33
    ("(3LOWER'_BOUNDS _:_./ _)" 10)
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    34
  "_LOWERS_U" :: "[pttrn, 'a => bool] => 'a set"           
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    35
    ("(3LOWER'_BOUNDS _./ _)" 10);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    36
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    37
translations
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    38
  "UPPER_BOUNDS x:A. P" == "UpperBounds A (Collect (\<lambda>x. P))"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    39
  "UPPER_BOUNDS x. P" == "UPPER_BOUNDS x:UNIV. P"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    40
  "LOWER_BOUNDS x:A. P" == "LowerBounds A (Collect (\<lambda>x. P))"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    41
  "LOWER_BOUNDS x. P" == "LOWER_BOUNDS x:UNIV. P";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    42
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    43
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    44
subsection {* Least and greatest elements *};
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    45
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    46
constdefs
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    47
  is_Least :: "('a::order) set => 'a => bool"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    48
  "is_Least B == is_LowerBound B B"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    49
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    50
  Least :: "('a::order) set => 'a"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    51
  "Least B == Eps (is_Least B)"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    52
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    53
  is_Greatest :: "('a::order) set => 'a => bool"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    54
  "is_Greatest B == is_UpperBound B B"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    55
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    56
  Greatest :: "('a::order) set => 'a" 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    57
  "Greatest B == Eps (is_Greatest B)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    58
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    59
syntax
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    60
  "_LEAST"    :: "[pttrn, 'a => bool] => 'a"  
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    61
    ("(3LLEAST _./ _)" 10)
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    62
  "_GREATEST" :: "[pttrn, 'a => bool] => 'a"  
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    63
    ("(3GREATEST _./ _)" 10);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    64
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    65
translations
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    66
  "LLEAST x. P" == "Least {x. P}"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    67
  "GREATEST x. P" == "Greatest {x. P}";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    68
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    69
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
    70
subsection {* Infimum and Supremum *};
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    71
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    72
text_raw {* \end{comment} *};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    73
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    74
text {* A supremum of an ordered set $B$ w.~r.~t.~$A$ 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    75
is defined as a least upperbound of $B$, which lies in $A$.
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    76
The definition of the supremum is based on the
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    77
existing definition (see HOL/Real/Lubs.thy).*};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    78
   
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    79
text{* If a supremum exists, then $\idt{Sup}\ap A\ap B$
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    80
is equal to the supremum. *};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    81
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    82
constdefs
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    83
  is_Sup :: "('a::order) set => 'a set => 'a => bool"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    84
  "is_Sup A B x == isLub A B x"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    85
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    86
  Sup :: "('a::order) set => 'a set => 'a"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    87
  "Sup A B == Eps (is_Sup A B)";
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    88
;
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    89
text_raw {* \begin{comment} *};
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    90
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    91
constdefs
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    92
  is_Inf :: "('a::order) set => 'a set => 'a => bool"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    93
  "is_Inf A B x == x:A & is_Greatest (LowerBounds A B) x"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    94
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    95
  Inf :: "('a::order) set => 'a set => 'a"
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    96
  "Inf A B == Eps (is_Inf A B)";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    97
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    98
syntax
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    99
  "_SUP" :: "[pttrn, 'a set, 'a => bool] => 'a set"
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   100
    ("(3SUP _:_./ _)" 10)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   101
  "_SUP_U" :: "[pttrn, 'a => bool] => 'a set"
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   102
    ("(3SUP _./ _)" 10)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   103
  "_INF" :: "[pttrn, 'a set, 'a => bool] => 'a set"
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   104
    ("(3INF _:_./ _)" 10)
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   105
  "_INF_U" :: "[pttrn, 'a => bool] => 'a set"
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   106
    ("(3INF _./ _)" 10);
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   107
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   108
translations
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   109
  "SUP x:A. P" == "Sup A (Collect (\<lambda>x. P))"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   110
  "SUP x. P" == "SUP x:UNIV. P"
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   111
  "INF x:A. P" == "Inf A (Collect (\<lambda>x. P))"
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   112
  "INF x. P" == "INF x:UNIV. P";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   113
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   114
text_raw {* \end{comment} *};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   115
;
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   116
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   117
text{* The supremum of $B$ is less than every upper bound
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   118
of $B$.*};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   119
7656
2f18c0ffc348 update from Gertrud;
wenzelm
parents: 7566
diff changeset
   120
lemma sup_le_ub: "isUb A B y ==> is_Sup A B s ==> s <= y";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   121
  by (unfold is_Sup_def, rule isLub_le_isUb);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   122
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   123
text {* The supremum $B$ is an upperbound for $B$. *};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   124
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   125
lemma sup_ub: "y:B ==> is_Sup A B s ==> y <= s";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   126
  by (unfold is_Sup_def, rule isLubD2);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   127
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   128
text{* The supremum of a non-empty set $B$ is greater
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   129
than a lower bound of $B$. *};
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   130
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   131
lemma sup_ub1: 
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
   132
  "[| ALL y:B. a <= y; is_Sup A B s; x:B |] ==> a <= s";
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   133
proof -; 
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   134
  assume "ALL y:B. a <= y" "is_Sup A B s" "x:B";
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   135
  have "a <= x"; by (rule bspec);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   136
  also; have "x <= s"; by (rule sup_ub);
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   137
  finally; show "a <= s"; .;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   138
qed;
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
   139
  
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
   140
end;