Removed uses of context element includes.
1 (* Title: HOL/Real/HahnBanach/Bounds.thy
3 Author: Gertrud Bauer, TU Munich
14 assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
15 and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
17 lemmas [elim?] = lub.least lub.upper
20 the_lub :: "'a::order set \<Rightarrow> 'a" where
21 "the_lub A = The (lub A)"
24 the_lub ("\<Squnion>_" [90] 90)
26 lemma the_lub_equality [elim?]:
28 shows "\<Squnion>A = (x::'a::order)"
30 interpret lub [A x] by fact
31 show ?thesis proof (unfold the_lub_def)
32 from `lub A x` show "The (lub A) = x"
34 fix x' assume lub': "lub A x'"
36 proof (rule order_antisym)
37 from lub' show "x' \<le> x"
39 fix a assume "a \<in> A"
40 then show "a \<le> x" ..
44 fix a assume "a \<in> A"
45 with lub' show "a \<le> x'" ..
53 assumes ex: "\<exists>x. lub A x"
54 shows "lub A (\<Squnion>A)"
56 from ex obtain x where x: "lub A x" ..
57 also from x have [symmetric]: "\<Squnion>A = x" ..
58 finally show ?thesis .
61 lemma lub_compat: "lub A x = isLub UNIV A x"
63 have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
64 by (rule ext) (simp only: isUb_def)
66 by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
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"
75 from ex_upper have "\<exists>y. isUb UNIV A y"
76 by (unfold isUb_def setle_def) blast
77 with nonempty have "\<exists>x. isLub UNIV A x"
78 by (rule reals_complete)
79 then show ?thesis by (simp only: lub_compat)