diff -r f8c2def19f84 -r fa56622bb7bc src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Aug 23 15:46:53 2011 -0700 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Aug 23 16:17:22 2011 -0700 @@ -921,7 +921,69 @@ thus ?lhs unfolding affine_dependent_explicit using assms by auto qed -subsection {* A general lemma. *} +subsection {* Connectedness of convex sets *} + +lemma connected_real_lemma: + fixes f :: "real \ 'a::metric_space" + assumes ab: "a \ b" and fa: "f a \ e1" and fb: "f b \ e2" + and dst: "\e x. a <= x \ x <= b \ 0 < e ==> \d > 0. \y. abs(y - x) < d \ dist(f y) (f x) < e" + and e1: "\y \ e1. \e > 0. \y'. dist y' y < e \ y' \ e1" + and e2: "\y \ e2. \e > 0. \y'. dist y' y < e \ y' \ e2" + and e12: "~(\x \ a. x <= b \ f x \ e1 \ f x \ e2)" + shows "\x \ a. x <= b \ f x \ e1 \ f x \ e2" (is "\ x. ?P x") +proof- + let ?S = "{c. \x \ a. x <= c \ f x \ e1}" + have Se: " \x. x \ ?S" apply (rule exI[where x=a]) by (auto simp add: fa) + have Sub: "\y. isUb UNIV ?S y" + apply (rule exI[where x= b]) + using ab fb e12 by (auto simp add: isUb_def setle_def) + from reals_complete[OF Se Sub] obtain l where + l: "isLub UNIV ?S l"by blast + have alb: "a \ l" "l \ b" using l ab fa fb e12 + apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + by (metis linorder_linear) + have ale1: "\z \ a. z < l \ f z \ e1" using l + apply (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + by (metis linorder_linear not_le) + have th1: "\z x e d :: real. z <= x + e \ e < d ==> z < x \ abs(z - x) < d" by arith + have th2: "\e x:: real. 0 < e ==> ~(x + e <= x)" by arith + have "\d::real. 0 < d \ 0 < d/2 \ d/2 < d" by simp + then have th3: "\d::real. d > 0 \ \e > 0. e < d" by blast + {assume le2: "f l \ e2" + from le2 fa fb e12 alb have la: "l \ a" by metis + hence lap: "l - a > 0" using alb by arith + from e2[rule_format, OF le2] obtain e where + e: "e > 0" "\y. dist y (f l) < e \ y \ e2" by metis + from dst[OF alb e(1)] obtain d where + d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis + let ?d' = "min (d/2) ((l - a)/2)" + have "?d' < d \ 0 < ?d' \ ?d' < l - a" using lap d(1) + by (simp add: min_max.less_infI2) + then have "\d'. d' < d \ d' >0 \ l - d' > a" by auto + then obtain d' where d': "d' > 0" "d' < d" "l - d' > a" by metis + from d e have th0: "\y. \y - l\ < d \ f y \ e2" by metis + from th0[rule_format, of "l - d'"] d' have "f (l - d') \ e2" by auto + moreover + have "f (l - d') \ e1" using ale1[rule_format, of "l -d'"] d' by auto + ultimately have False using e12 alb d' by auto} + moreover + {assume le1: "f l \ e1" + from le1 fa fb e12 alb have lb: "l \ b" by metis + hence blp: "b - l > 0" using alb by arith + from e1[rule_format, OF le1] obtain e where + e: "e > 0" "\y. dist y (f l) < e \ y \ e1" by metis + from dst[OF alb e(1)] obtain d where + d: "d > 0" "\y. \y - l\ < d \ dist (f y) (f l) < e" by metis + have "\d::real. 0 < d \ d/2 < d \ 0 < d/2" by simp + then have "\d'. d' < d \ d' >0" using d(1) by blast + then obtain d' where d': "d' > 0" "d' < d" by metis + from d e have th0: "\y. \y - l\ < d \ f y \ e1" by auto + hence "\y. l \ y \ y \ l + d' \ f y \ e1" using d' by auto + with ale1 have "\y. a \ y \ y \ l + d' \ f y \ e1" by auto + with l d' have False + by (auto simp add: isLub_def isUb_def setle_def setge_def leastP_def) } + ultimately show ?thesis using alb by metis +qed lemma convex_connected: fixes s :: "'a::real_normed_vector set"