(* Title: HOL/Analysis/Convex_Euclidean_Space.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen
*)
section \<open>Convex Sets and Functions on (Normed) Euclidean Spaces\<close>
theory Convex_Euclidean_Space imports
Convex
Topology_Euclidean_Space begin
subsection\<^marker>\<open>tag unimportant\<close> \<open>Topological Properties of Convex Sets and Functions\<close>
lemma aff_dim_cball: fixes a :: "'n::euclidean_space" assumes"e > 0" shows"aff_dim (cball a e) = int (DIM('n))" proof - have"(\x. a + x) ` (cball 0 e) \ cball a e" unfolding cball_def dist_norm by auto thenhave"aff_dim (cball (0 :: 'n::euclidean_space) e) \ aff_dim (cball a e)" using aff_dim_translation_eq[of a "cball 0 e"]
aff_dim_subset[of "(+) a ` cball 0 e""cball a e"] by auto moreoverhave"aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" using hull_inc[of "(0 :: 'n::euclidean_space)""cball 0 e"]
centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) ultimatelyshow ?thesis using aff_dim_le_DIM[of "cball a e"] by auto qed
lemma aff_dim_open: fixes S :: "'n::euclidean_space set" assumes"open S" and"S \ {}" shows"aff_dim S = int (DIM('n))" proof - obtain x where"x \ S" using assms by auto thenobtain e where e: "e > 0""cball x e \ S" using open_contains_cball[of S] assms by auto thenhave"aff_dim (cball x e) \ aff_dim S" using aff_dim_subset by auto with e show ?thesis using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto qed
lemma low_dim_interior: fixes S :: "'n::euclidean_space set" assumes"\ aff_dim S = int (DIM('n))" shows"interior S = {}" proof - have"aff_dim(interior S) \ aff_dim S" using interior_subset aff_dim_subset[of "interior S" S] by auto thenshow ?thesis using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto qed
corollary empty_interior_lowdim: fixes S :: "'n::euclidean_space set" shows"dim S < DIM ('n) \ interior S = {}" by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV)
corollary aff_dim_nonempty_interior: fixes S :: "'a::euclidean_space set" shows"interior S \ {} \ aff_dim S = DIM('a)" by (metis low_dim_interior)
subsection \<open>Relative interior of a set\<close>
definition\<^marker>\<open>tag important\<close> "rel_interior S =
{x. \<exists>T. openin (top_of_set (affine hull S)) T \<and> x \<in> T \<and> T \<subseteq> S}"
lemma rel_interior_maximal: "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)" by (auto simp: rel_interior_def)
lemma rel_interior: "rel_interior S = {x \ S. \T. open T \ x \ T \ T \ affine hull S \ S}"
(is"?lhs = ?rhs") proof show"?lhs \ ?rhs" by (force simp add: rel_interior_def openin_open)
{ fix x T assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S" thenhave **: "x \ T \ affine hull S" using hull_inc by auto with * have"\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S" by (rule_tac x = "T \ (affine hull S)" in exI) auto
} thenshow"?rhs \ ?lhs" by (force simp add: rel_interior_def openin_open) qed
lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)" by (auto simp: rel_interior)
lemma mem_rel_interior_ball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)"
(is"?lhs = ?rhs") proof assume ?rhs thenshow ?lhs by (simp add: rel_interior) (meson Elementary_Metric_Spaces.open_ball centre_in_ball) qed (force simp: rel_interior open_contains_ball)
lemma rel_interior_ball: "rel_interior S = {x \ S. \e. e > 0 \ ball x e \ affine hull S \ S}" using mem_rel_interior_ball [of _ S] by auto
lemma mem_rel_interior_cball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)"
(is"?lhs = ?rhs") proof assume ?rhs thenobtain e where"x \ S" "e > 0" "cball x e \ affine hull S \ S" by (auto simp: rel_interior) thenhave"ball x e \ affine hull S \ S" by auto thenshow ?lhs using\<open>0 < e\<close> \<open>x \<in> S\<close> rel_interior_ball by auto qed (force simp: rel_interior open_contains_cball)
lemma rel_interior_cball: "rel_interior S = {x \ S. \e. e > 0 \ cball x e \ affine hull S \ S}" using mem_rel_interior_cball [of _ S] by auto
lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space"shows"rel_interior {a} = {a}" proof - have"\x::real. 0 < x" using zero_less_one by blast thenshow ?thesis by (auto simp: rel_interior_ball) qed
lemma subset_rel_interior: fixes S T :: "'n::euclidean_space set" assumes"S \ T" and"affine hull S = affine hull T" shows"rel_interior S \ rel_interior T" using assms by (auto simp: rel_interior_def)
lemma rel_interior_subset: "rel_interior S \ S" by (auto simp: rel_interior_def)
lemma rel_interior_subset_closure: "rel_interior S \ closure S" using rel_interior_subset by (auto simp: closure_def)
lemma interior_subset_rel_interior: "interior S \ rel_interior S" by (auto simp: rel_interior interior_def)
lemma interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes"aff_dim S = int(DIM('n))" shows"rel_interior S = interior S" proof - have"affine hull S = UNIV" using assms affine_hull_UNIV[of S] by auto thenshow ?thesis unfolding rel_interior interior_def by auto qed
lemma rel_interior_interior: fixes S :: "'n::euclidean_space set" assumes"affine hull S = UNIV" shows"rel_interior S = interior S" using assms unfolding rel_interior interior_def by auto
lemma rel_interior_open: fixes S :: "'n::euclidean_space set" assumes"open S" shows"rel_interior S = S" by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset)
lemma interior_rel_interior_gen: fixes S :: "'n::euclidean_space set" shows"interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" by (metis interior_rel_interior low_dim_interior)
lemma rel_interior_nonempty_interior: fixes S :: "'n::euclidean_space set" shows"interior S \ {} \ rel_interior S = interior S" by (metis interior_rel_interior_gen)
lemma affine_hull_nonempty_interior: fixes S :: "'n::euclidean_space set" shows"interior S \ {} \ affine hull S = UNIV" by (metis affine_hull_UNIV interior_rel_interior_gen)
lemma rel_interior_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows"rel_interior (affine hull S) = affine hull S" proof - have *: "rel_interior (affine hull S) \ affine hull S" using rel_interior_subset by auto
{ fix x assume x: "x \ affine hull S"
define e :: real where"e = 1" thenhave"e > 0""ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto thenhave"x \ rel_interior (affine hull S)" using x rel_interior_ball[of "affine hull S"] by auto
} thenshow ?thesis using * by auto qed
lemma rel_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes"convex S" and"c \ rel_interior S" and"x \ S" and"0 < e" and"e \ 1" shows"x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where"d > 0"and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto
{ fix y assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S" have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using\<open>e > 0\<close> by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have"x \ affine hull S" using assms hull_subset[of S] by auto moreoverhave"1 / e + - ((1 - e) / e) = 1" using\<open>e > 0\<close> left_diff_distrib[of "1" "(1-e)" "1/e"] by auto ultimatelyhave **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ affine hull S" using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)""-((1 - e) / e)"] by (simp add: algebra_simps) have"c - ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = (1 / e) *\<^sub>R (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" using\<open>e > 0\<close> by (auto simp: euclidean_eq_iff[where'a='a] field_simps inner_simps) thenhave"dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm norm_scaleR[symmetric] by auto alsohave"\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) alsohave"\ < d" using as[unfolded dist_norm] and\<open>e > 0\<close> by (auto simp:pos_divide_less_eq[OF \<open>e > 0\<close>] mult.commute) finallyhave"(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ S" using"**" d by auto thenhave"y \ S" using * convexD [OF \<open>convex S\<close>] assms(3-5) by (metis diff_add_cancel diff_ge_0_iff_ge le_add_same_cancel1 less_eq_real_def)
} thenhave"ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S" by auto moreoverhave"e * d > 0" using\<open>e > 0\<close> \<open>d > 0\<close> by simp moreoverhave c: "c \ S" using assms rel_interior_subset by auto moreoverfrom c have"x - e *\<^sub>R (x - c) \ S" using convexD_alt[of S x c e] assms by (metis diff_add_eq diff_diff_eq2 less_eq_real_def scaleR_diff_left scaleR_one scale_right_diff_distrib) ultimatelyshow ?thesis using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \e > 0\ by auto qed
lemma interior_real_atLeast [simp]: fixes a :: real shows"interior {a..} = {a<..}" proof -
{ fix y have"ball y (y - a) \ {a..}" by (auto simp: dist_norm) moreoverassume"a < y" ultimatelyhave"y \ interior {a..}" by (force simp add: mem_interior)
} moreover
{ fix y assume"y \ interior {a..}" thenobtain e where e: "e > 0""cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreoverfrom e have"y - e \ cball y e" by (auto simp: cball_def dist_norm) ultimatelyhave"a \ y - e" by blast thenhave"a < y"using e by auto
} ultimatelyshow ?thesis by auto qed
lemma continuous_ge_on_Ioo: assumes"continuous_on {c..d} g""\x. x \ {c<.. g x \ a" "c < d" "x \ {c..d}" shows"g (x::real) \ (a::real)"
proof- from assms(3) have"{c..d} = closure {c<..by (rule closure_greaterThanLessThan[symmetric]) alsofrom assms(2) have"{c<.. (g -` {a..} \ {c..d})" by auto hence"closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono) alsofrom assms(1) have"closed (g -` {a..} \ {c..d})" by (auto simp: continuous_on_closed_vimage) hence"closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp finallyshow ?thesis using\<open>x \<in> {c..d}\<close> by auto qed