author | wenzelm |
Fri, 13 May 2011 23:58:40 +0200 | |
changeset 42795 | 66fcc9882784 |
parent 41413 | 64cd30d6b0b8 |
child 44887 | 7ca82df6e951 |
permissions | -rw-r--r-- |
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
1 |
(* Title: HOL/Hahn_Banach/Bounds.thy |
7566 | 2 |
Author: Gertrud Bauer, TU Munich |
3 |
*) |
|
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
4 |
|
9035 | 5 |
header {* Bounds *} |
7808 | 6 |
|
25596 | 7 |
theory Bounds |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
32960
diff
changeset
|
8 |
imports Main "~~/src/HOL/Library/ContNotDenum" |
25596 | 9 |
begin |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
10 |
|
13515 | 11 |
locale lub = |
12 |
fixes A and x |
|
13 |
assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b" |
|
14 |
and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x" |
|
15 |
||
16 |
lemmas [elim?] = lub.least lub.upper |
|
17 |
||
19736 | 18 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
19 |
the_lub :: "'a::order set \<Rightarrow> 'a" where |
19736 | 20 |
"the_lub A = The (lub A)" |
14653 | 21 |
|
21210 | 22 |
notation (xsymbols) |
19736 | 23 |
the_lub ("\<Squnion>_" [90] 90) |
7535
599d3414b51d
The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff
changeset
|
24 |
|
13515 | 25 |
lemma the_lub_equality [elim?]: |
27611 | 26 |
assumes "lub A x" |
13515 | 27 |
shows "\<Squnion>A = (x::'a::order)" |
27611 | 28 |
proof - |
29234 | 29 |
interpret lub A x by fact |
27612 | 30 |
show ?thesis |
31 |
proof (unfold the_lub_def) |
|
27611 | 32 |
from `lub A x` show "The (lub A) = x" |
33 |
proof |
|
34 |
fix x' assume lub': "lub A x'" |
|
35 |
show "x' = x" |
|
36 |
proof (rule order_antisym) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
37 |
from lub' show "x' \<le> x" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
38 |
proof |
27611 | 39 |
fix a assume "a \<in> A" |
40 |
then show "a \<le> x" .. |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
41 |
qed |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
42 |
show "x \<le> x'" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
43 |
proof |
27611 | 44 |
fix a assume "a \<in> A" |
45 |
with lub' show "a \<le> x'" .. |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
46 |
qed |
13515 | 47 |
qed |
48 |
qed |
|
49 |
qed |
|
50 |
qed |
|
7917 | 51 |
|
13515 | 52 |
lemma the_lubI_ex: |
53 |
assumes ex: "\<exists>x. lub A x" |
|
54 |
shows "lub A (\<Squnion>A)" |
|
55 |
proof - |
|
56 |
from ex obtain x where x: "lub A x" .. |
|
57 |
also from x have [symmetric]: "\<Squnion>A = x" .. |
|
58 |
finally show ?thesis . |
|
59 |
qed |
|
7917 | 60 |
|
13515 | 61 |
lemma lub_compat: "lub A x = isLub UNIV A x" |
62 |
proof - |
|
63 |
have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)" |
|
64 |
by (rule ext) (simp only: isUb_def) |
|
65 |
then show ?thesis |
|
66 |
by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast |
|
9035 | 67 |
qed |
13515 | 68 |
|
69 |
lemma real_complete: |
|
70 |
fixes A :: "real set" |
|
71 |
assumes nonempty: "\<exists>a. a \<in> A" |
|
72 |
and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y" |
|
73 |
shows "\<exists>x. lub A x" |
|
74 |
proof - |
|
75 |
from ex_upper have "\<exists>y. isUb UNIV A y" |
|
27612 | 76 |
unfolding isUb_def setle_def by blast |
13515 | 77 |
with nonempty have "\<exists>x. isLub UNIV A x" |
78 |
by (rule reals_complete) |
|
79 |
then show ?thesis by (simp only: lub_compat) |
|
80 |
qed |
|
81 |
||
10687 | 82 |
end |