author | paulson <lp15@cam.ac.uk> |
Fri, 30 Dec 2022 17:48:41 +0000 | |
changeset 76832 | ab08604729a2 |
parent 66453 | cc19f7ca2ed6 |
child 80914 | d97fdabd9e2b |
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 |
|
58889 | 5 |
section \<open>Bounds\<close> |
7808 | 6 |
|
25596 | 7 |
theory Bounds |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63970
diff
changeset
|
8 |
imports Main "HOL-Analysis.Continuum_Not_Denumerable" |
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 |
||
58745 | 18 |
definition the_lub :: "'a::order set \<Rightarrow> 'a" ("\<Squnion>_" [90] 90) |
44887 | 19 |
where "the_lub A = The (lub A)" |
14653 | 20 |
|
13515 | 21 |
lemma the_lub_equality [elim?]: |
27611 | 22 |
assumes "lub A x" |
13515 | 23 |
shows "\<Squnion>A = (x::'a::order)" |
27611 | 24 |
proof - |
29234 | 25 |
interpret lub A x by fact |
27612 | 26 |
show ?thesis |
27 |
proof (unfold the_lub_def) |
|
58744 | 28 |
from \<open>lub A x\<close> show "The (lub A) = x" |
27611 | 29 |
proof |
30 |
fix x' assume lub': "lub A x'" |
|
31 |
show "x' = x" |
|
32 |
proof (rule order_antisym) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
33 |
from lub' show "x' \<le> x" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
34 |
proof |
27611 | 35 |
fix a assume "a \<in> A" |
36 |
then show "a \<le> x" .. |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
37 |
qed |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
38 |
show "x \<le> x'" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
39 |
proof |
27611 | 40 |
fix a assume "a \<in> A" |
41 |
with lub' show "a \<le> x'" .. |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
31795
diff
changeset
|
42 |
qed |
13515 | 43 |
qed |
44 |
qed |
|
45 |
qed |
|
46 |
qed |
|
7917 | 47 |
|
13515 | 48 |
lemma the_lubI_ex: |
49 |
assumes ex: "\<exists>x. lub A x" |
|
50 |
shows "lub A (\<Squnion>A)" |
|
51 |
proof - |
|
52 |
from ex obtain x where x: "lub A x" .. |
|
53 |
also from x have [symmetric]: "\<Squnion>A = x" .. |
|
54 |
finally show ?thesis . |
|
55 |
qed |
|
7917 | 56 |
|
54263
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
44887
diff
changeset
|
57 |
lemma real_complete: "\<exists>a::real. a \<in> A \<Longrightarrow> \<exists>y. \<forall>a \<in> A. a \<le> y \<Longrightarrow> \<exists>x. lub A x" |
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents:
44887
diff
changeset
|
58 |
by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def) |
13515 | 59 |
|
10687 | 60 |
end |