wenzelm@31795: (* Title: HOL/Hahn_Banach/Bounds.thy wenzelm@7566: Author: Gertrud Bauer, TU Munich wenzelm@7566: *) wenzelm@7535: wenzelm@9035: header {* Bounds *} wenzelm@7808: haftmann@25596: theory Bounds wenzelm@41413: imports Main "~~/src/HOL/Library/ContNotDenum" haftmann@25596: begin wenzelm@7535: wenzelm@13515: locale lub = wenzelm@13515: fixes A and x wenzelm@13515: assumes least [intro?]: "(\a. a \ A \ a \ b) \ x \ b" wenzelm@13515: and upper [intro?]: "a \ A \ a \ x" wenzelm@13515: wenzelm@13515: lemmas [elim?] = lub.least lub.upper wenzelm@13515: wenzelm@44887: definition the_lub :: "'a::order set \ 'a" wenzelm@44887: where "the_lub A = The (lub A)" wenzelm@14653: wenzelm@21210: notation (xsymbols) wenzelm@19736: the_lub ("\_" [90] 90) wenzelm@7535: wenzelm@13515: lemma the_lub_equality [elim?]: ballarin@27611: assumes "lub A x" wenzelm@13515: shows "\A = (x::'a::order)" ballarin@27611: proof - ballarin@29234: interpret lub A x by fact wenzelm@27612: show ?thesis wenzelm@27612: proof (unfold the_lub_def) ballarin@27611: from `lub A x` show "The (lub A) = x" ballarin@27611: proof ballarin@27611: fix x' assume lub': "lub A x'" ballarin@27611: show "x' = x" ballarin@27611: proof (rule order_antisym) wenzelm@32960: from lub' show "x' \ x" wenzelm@32960: proof ballarin@27611: fix a assume "a \ A" ballarin@27611: then show "a \ x" .. wenzelm@32960: qed wenzelm@32960: show "x \ x'" wenzelm@32960: proof ballarin@27611: fix a assume "a \ A" ballarin@27611: with lub' show "a \ x'" .. wenzelm@32960: qed wenzelm@13515: qed wenzelm@13515: qed wenzelm@13515: qed wenzelm@13515: qed wenzelm@7917: wenzelm@13515: lemma the_lubI_ex: wenzelm@13515: assumes ex: "\x. lub A x" wenzelm@13515: shows "lub A (\A)" wenzelm@13515: proof - wenzelm@13515: from ex obtain x where x: "lub A x" .. wenzelm@13515: also from x have [symmetric]: "\A = x" .. wenzelm@13515: finally show ?thesis . wenzelm@13515: qed wenzelm@7917: wenzelm@13515: lemma lub_compat: "lub A x = isLub UNIV A x" wenzelm@13515: proof - wenzelm@13515: have "isUb UNIV A = (\x. A *<= x \ x \ UNIV)" wenzelm@13515: by (rule ext) (simp only: isUb_def) wenzelm@13515: then show ?thesis wenzelm@13515: by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast wenzelm@9035: qed wenzelm@13515: wenzelm@13515: lemma real_complete: wenzelm@13515: fixes A :: "real set" wenzelm@13515: assumes nonempty: "\a. a \ A" wenzelm@13515: and ex_upper: "\y. \a \ A. a \ y" wenzelm@13515: shows "\x. lub A x" wenzelm@13515: proof - wenzelm@13515: from ex_upper have "\y. isUb UNIV A y" wenzelm@27612: unfolding isUb_def setle_def by blast wenzelm@13515: with nonempty have "\x. isLub UNIV A x" wenzelm@13515: by (rule reals_complete) wenzelm@13515: then show ?thesis by (simp only: lub_compat) wenzelm@13515: qed wenzelm@13515: wenzelm@10687: end