diff --git a/src/HOL/Analysis/Convex_Euclidean_Space.thy b/src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy @@ -1,2524 +1,3129 @@ (* 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 \Convex Sets and Functions on (Normed) Euclidean Spaces\ theory Convex_Euclidean_Space imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets and Functions\ lemma convex_supp_sum: assumes "convex S" and 1: "supp_sum u I = 1" and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" shows "supp_sum (\i. u i *\<^sub>R f i) I \ S" proof - have fin: "finite {i \ I. u i \ 0}" using 1 sum.infinite by (force simp: supp_sum_def support_on_def) then have eq: "supp_sum (\i. u i *\<^sub>R f i) I = sum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) show ?thesis apply (simp add: eq) apply (rule convex_sum [OF fin \convex S\]) using 1 assms apply (auto simp: supp_sum_def support_on_def) done qed lemma closure_bounded_linear_image_subset: assumes f: "bounded_linear f" shows "f ` closure S \ closure (f ` S)" using linear_continuous_on [OF f] closed_closure closure_subset by (rule image_closure_subset) lemma closure_linear_image_subset: fixes f :: "'m::euclidean_space \ 'n::real_normed_vector" assumes "linear f" shows "f ` (closure S) \ closure (f ` S)" using assms unfolding linear_conv_bounded_linear by (rule closure_bounded_linear_image_subset) lemma closed_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "closed S" and f: "linear f" "inj f" shows "closed (f ` S)" proof - obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF f] by blast then have confg: "continuous_on (range f) g" using linear_continuous_on linear_conv_bounded_linear by blast have [simp]: "g ` f ` S = S" using g by (simp add: image_comp) have cgf: "closed (g ` f ` S)" by (simp add: \g \ f = id\ S image_comp) have [simp]: "(range f \ g -` S) = f ` S" using g unfolding o_def id_def image_def by auto metis+ show ?thesis proof (rule closedin_closed_trans [of "range f"]) show "closedin (top_of_set (range f)) (f ` S)" using continuous_closedin_preimage [OF confg cgf] by simp show "closed (range f)" apply (rule closed_injective_image_subspace) using f apply (auto simp: linear_linear linear_injective_0) done qed qed lemma closed_injective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" shows "(closed(image f s) \ closed s)" by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) lemma closure_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; inj f\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym) apply (simp add: closure_linear_image_subset) by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) lemma closure_bounded_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; bounded S\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym, simp add: closure_linear_image_subset) apply (rule closure_minimal, simp add: closure_subset image_mono) by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) lemma closure_scaleR: fixes S :: "'a::real_normed_vector set" shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" proof show "((*\<^sub>R) c) ` (closure S) \ closure (((*\<^sub>R) c) ` S)" using bounded_linear_scaleR_right by (rule closure_bounded_linear_image_subset) show "closure (((*\<^sub>R) c) ` S) \ ((*\<^sub>R) c) ` (closure S)" by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) qed lemma sphere_eq_empty [simp]: fixes a :: "'a::{real_normed_vector, perfect_space}" shows "sphere a r = {} \ r < 0" by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) lemma cone_closure: fixes S :: "'a::real_normed_vector set" assumes "cone S" shows "cone (closure S)" proof (cases "S = {}") case True then show ?thesis by auto next case False then have "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have "0 \ closure S \ (\c. c > 0 \ (*\<^sub>R) c ` closure S = closure S)" using closure_subset by (auto simp: closure_scaleR) then show ?thesis using False cone_iff[of "closure S"] by auto qed corollary component_complement_connected: fixes S :: "'a::real_normed_vector set" assumes "connected S" "C \ components (-S)" shows "connected(-C)" using component_diff_connected [of S UNIV] assms by (auto simp: Compl_eq_Diff_UNIV) proposition clopen: fixes S :: "'a :: real_normed_vector set" shows "closed S \ open S \ S = {} \ S = UNIV" by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format]) corollary compact_open: fixes S :: "'a :: euclidean_space set" shows "compact S \ open S \ S = {}" by (auto simp: compact_eq_bounded_closed clopen) corollary finite_imp_not_open: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\finite S; open S\ \ S={}" using clopen [of S] finite_imp_closed not_bounded_UNIV by blast corollary empty_interior_finite: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "finite S \ interior S = {}" by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) text \Balls, being convex, are connected.\ lemma convex_local_global_minimum: fixes s :: "'a::real_normed_vector set" assumes "e > 0" and "convex_on s f" and "ball x e \ s" and "\y\ball x e. f x \ f y" shows "\y\s. f x \ f y" proof (rule ccontr) have "x \ s" using assms(1,3) by auto assume "\ ?thesis" then obtain y where "y\s" and y: "f x > f y" by auto then have xy: "0 < dist x y" by auto then obtain u where "0 < u" "u \ 1" and u: "u < e / dist x y" using field_lbound_gt_zero[of 1 "e / dist x y"] xy \e>0\ by auto then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \ (1-u) * f x + u * f y" using \x\s\ \y\s\ using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto moreover have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps) have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF \0] unfolding dist_norm[symmetric] using u unfolding pos_less_divide_eq[OF xy] by auto then have "f x \ f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" using assms(4) by auto ultimately show False using mult_strict_left_mono[OF y \u>0\] unfolding left_diff_distrib by auto qed lemma convex_ball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (ball x e)" proof (auto simp: convex_def) fix y z assume yz: "dist x y < e" "dist x z < e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" using convex_bound_lt[OF yz uv] by auto qed lemma convex_cball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (cball x e)" proof - { fix y z assume yz: "dist x y \ e" "dist x z \ e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ e" using convex_bound_le[OF yz uv] by auto } then show ?thesis by (auto simp: convex_def Ball_def) qed lemma connected_ball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (ball x e)" using convex_connected convex_ball by auto lemma connected_cball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (cball x e)" using convex_connected convex_cball by auto lemma bounded_convex_hull: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "bounded (convex hull s)" proof - from assms obtain B where B: "\x\s. norm x \ B" unfolding bounded_iff by auto show ?thesis apply (rule bounded_subset[OF bounded_cball, of _ 0 B]) unfolding subset_hull[of convex, OF convex_cball] unfolding subset_eq mem_cball dist_norm using B apply auto done qed lemma finite_imp_bounded_convex_hull: fixes s :: "'a::real_normed_vector set" shows "finite s \ bounded (convex hull s)" using bounded_convex_hull finite_imp_bounded by auto 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 then have "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 moreover have "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"]) ultimately show ?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 then obtain e where e: "e > 0" "cball x e \ S" using open_contains_cball[of S] assms by auto then have "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 then show ?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 \Relative interior of a set\ definition\<^marker>\tag important\ "rel_interior S = {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: "\S \ T; affine hull S = affine hull T\ \ (rel_interior S) \ (rel_interior T)" by (auto simp: rel_interior_def) 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}" unfolding rel_interior_def[of S] openin_open[of "affine hull S"] apply auto proof - fix x T assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S" then have **: "x \ T \ affine hull S" using hull_inc by auto show "\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S" apply (rule_tac x = "T \ (affine hull S)" in exI) using * ** apply auto done 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)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_ball) apply (rule_tac x = "ball x e" in exI, simp) done 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)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_cball) apply (rule_tac x = "ball x e" in exI) apply (simp add: subset_trans [OF ball_subset_cball], auto) done 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_empty [simp]: "rel_interior {} = {}" by (auto simp: rel_interior_def) lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" by (metis affine_hull_eq affine_sing) lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" apply (auto simp: rel_interior_ball) apply (rule_tac x=1 in exI, force) done 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 then show ?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_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) 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" then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto then have "x \ rel_interior (affine hull S)" using x rel_interior_ball[of "affine hull S"] by auto } then show ?thesis using * by auto qed lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" by (metis open_UNIV rel_interior_open) 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 \e > 0\ by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "x \ affine hull S" using assms hull_subset[of S] by auto moreover have "1 / e + - ((1 - e) / e) = 1" using \e > 0\ left_diff_distrib[of "1" "(1-e)" "1/e"] by auto ultimately have **: "(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 "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] apply (rule arg_cong[where f=norm]) using \e > 0\ apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) done also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[THEN subsetD]) unfolding mem_ball using assms(3-5) ** apply auto done } then have "ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S" by auto moreover have "e * d > 0" using \e > 0\ \d > 0\ by simp moreover have c: "c \ S" using assms rel_interior_subset by auto moreover from c have "x - e *\<^sub>R (x - c) \ S" using convexD_alt[of S x c e] apply (simp add: algebra_simps) using assms apply auto done ultimately show ?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 assume "a < y" then have "y \ interior {a..}" apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {a..}" then obtain e where e: "e > 0" "cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreover from e have "y - e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y - e" by blast then have "a < y" using e by auto } ultimately show ?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<.. (g -` {a..} \ {c..d})" by auto hence "closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono) also from 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 finally show ?thesis using \x \ {c..d}\ by auto qed lemma interior_real_atMost [simp]: fixes a :: real shows "interior {..a} = {.. y" then have "y \ interior {..a}" apply (simp add: mem_interior) apply (rule_tac x="(a-y)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {..a}" then obtain e where e: "e > 0" "cball y e \ {..a}" using mem_interior_cball[of y "{..a}"] by auto moreover from e have "y + e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y + e" by auto then have "a > y" using e by auto } ultimately show ?thesis by auto qed lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<.. {..b}" by auto also have "interior \ = {a<..} \ {.. = {a<.. {}" using assms unfolding set_eq_iff by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) then show ?thesis using interior_rel_interior_gen[of "cbox a b", symmetric] by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) qed lemma rel_interior_real_semiline [simp]: fixes a :: real shows "rel_interior {a..} = {a<..}" proof - have *: "{a<..} \ {}" unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"] by (auto split: if_split_asm) qed subsubsection \Relative open sets\ definition\<^marker>\tag important\ "rel_open S \ rel_interior S = S" lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" unfolding rel_open_def rel_interior_def apply auto using openin_subopen[of "top_of_set (affine hull S)" S] apply auto done lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" apply (simp add: rel_interior_def) apply (subst openin_subopen, blast) done lemma openin_set_rel_interior: "openin (top_of_set S) (rel_interior S)" by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) lemma affine_rel_open: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "rel_open S" unfolding rel_open_def using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] by metis lemma affine_closed: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "closed S" proof - { assume "S \ {}" then obtain L where L: "subspace L" "affine_parallel S L" using assms affine_parallel_subspace[of S] by auto then obtain a where a: "S = ((+) a ` L)" using affine_parallel_def[of L S] affine_parallel_commut by auto from L have "closed L" using closed_subspace by auto then have "closed S" using closed_translation a by auto } then show ?thesis by auto qed lemma closure_affine_hull: fixes S :: "'n::euclidean_space set" shows "closure S \ affine hull S" by (intro closure_minimal hull_subset affine_closed affine_affine_hull) lemma closure_same_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "affine hull (closure S) = affine hull S" proof - have "affine hull (closure S) \ affine hull S" using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto moreover have "affine hull (closure S) \ affine hull S" using hull_mono[of "S" "closure S" "affine"] closure_subset by auto ultimately show ?thesis by auto qed lemma closure_aff_dim [simp]: fixes S :: "'n::euclidean_space set" shows "aff_dim (closure S) = aff_dim S" proof - have "aff_dim S \ aff_dim (closure S)" using aff_dim_subset closure_subset by auto moreover have "aff_dim (closure S) \ aff_dim (affine hull S)" using aff_dim_subset closure_affine_hull by blast moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto ultimately show ?thesis by auto qed lemma rel_interior_closure_convex_shrink: fixes S :: "_::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ closure S" and "e > 0" 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 have "\y \ S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x], auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have zball: "z \ ball c d" using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (simp add: norm_minus_commute) (simp add: field_simps) have "x \ affine hull S" using closure_affine_hull assms by auto moreover have "y \ affine hull S" using \y \ S\ hull_subset[of S] by auto moreover have "c \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto ultimately have "z \ affine hull S" using z_def affine_affine_hull[of S] mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] assms by simp then have "z \ S" using d zball by auto obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d" using zball open_ball[of c d] openE[of "ball c d" z] by auto then have "ball z d1 \ affine hull S \ ball c d \ affine hull S" by auto then have "ball z d1 \ affine hull S \ S" using d by auto then have "z \ rel_interior S" using mem_rel_interior_ball using \d1 > 0\ \z \ S\ by auto then have "y - e *\<^sub>R (y - z) \ rel_interior S" using rel_interior_convex_shrink[of S z y e] assms \y \ S\ by auto then show ?thesis using * by auto qed lemma rel_interior_eq: "rel_interior s = s \ openin(top_of_set (affine hull s)) s" using rel_open rel_open_def by blast lemma rel_interior_openin: "openin(top_of_set (affine hull s)) s \ rel_interior s = s" by (simp add: rel_interior_eq) lemma rel_interior_affine: fixes S :: "'n::euclidean_space set" shows "affine S \ rel_interior S = S" using affine_rel_open rel_open_def by auto lemma rel_interior_eq_closure: fixes S :: "'n::euclidean_space set" shows "rel_interior S = closure S \ affine S" proof (cases "S = {}") case True then show ?thesis by auto next case False show ?thesis proof assume eq: "rel_interior S = closure S" have "S = {} \ S = affine hull S" apply (rule connected_clopen [THEN iffD1, rule_format]) apply (simp add: affine_imp_convex convex_connected) apply (rule conjI) apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) done with False have "affine hull S = S" by auto then show "affine S" by (metis affine_hull_eq) next assume "affine S" then show "rel_interior S = closure S" by (simp add: rel_interior_affine affine_closed) qed qed subsubsection\<^marker>\tag unimportant\\Relative interior preserves under linear transformations\ lemma rel_interior_translation_aux: fixes a :: "'n::euclidean_space" shows "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" proof - { fix x assume x: "x \ rel_interior S" then obtain T where "open T" "x \ T \ S" "T \ affine hull S \ S" using mem_rel_interior[of x S] by auto then have "open ((\x. a + x) ` T)" and "a + x \ ((\x. a + x) ` T) \ ((\x. a + x) ` S)" and "((\x. a + x) ` T) \ affine hull ((\x. a + x) ` S) \ (\x. a + x) ` S" using affine_hull_translation[of a S] open_translation[of T a] x by auto then have "a + x \ rel_interior ((\x. a + x) ` S)" using mem_rel_interior[of "a+x" "((\x. a + x) ` S)"] by auto } then show ?thesis by auto qed lemma rel_interior_translation: fixes a :: "'n::euclidean_space" shows "rel_interior ((\x. a + x) ` S) = (\x. a + x) ` rel_interior S" proof - have "(\x. (-a) + x) ` rel_interior ((\x. a + x) ` S) \ rel_interior S" using rel_interior_translation_aux[of "-a" "(\x. a + x) ` S"] translation_assoc[of "-a" "a"] by auto then have "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] by auto then show ?thesis using rel_interior_translation_aux[of a S] by auto qed lemma affine_hull_linear_image: assumes "bounded_linear f" shows "f ` (affine hull s) = affine hull f ` s" proof - interpret f: bounded_linear f by fact have "affine {x. f x \ affine hull f ` s}" unfolding affine_def by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) moreover have "affine {x. x \ f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) ultimately show ?thesis by (auto simp: hull_inc elim!: hull_induct) qed lemma rel_interior_injective_on_span_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and S :: "'m::euclidean_space set" assumes "bounded_linear f" and "inj_on f (span S)" shows "rel_interior (f ` S) = f ` (rel_interior S)" proof - { fix z assume z: "z \ rel_interior (f ` S)" then have "z \ f ` S" using rel_interior_subset[of "f ` S"] by auto then obtain x where x: "x \ S" "f x = z" by auto obtain e2 where e2: "e2 > 0" "cball z e2 \ affine hull (f ` S) \ (f ` S)" using z rel_interior_cball[of "f ` S"] by auto obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K" using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto define e1 where "e1 = 1 / K" then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x" using K pos_le_divide_eq[of e1] by auto define e where "e = e1 * e2" then have "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball x e \ affine hull S" then have h1: "f y \ affine hull (f ` S)" using affine_hull_linear_image[of f S] assms by auto from y have "norm (x-y) \ e1 * e2" using cball_def[of x e] dist_norm[of x y] e_def by auto moreover have "f x - f y = f (x - y)" using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto moreover have "e1 * norm (f (x-y)) \ norm (x - y)" using e1 by auto ultimately have "e1 * norm ((f x)-(f y)) \ e1 * e2" by auto then have "f y \ cball z e2" using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto then have "f y \ f ` S" using y e2 h1 by auto then have "y \ S" using assms y hull_subset[of S] affine_hull_subset_span inj_on_image_mem_iff [OF \inj_on f (span S)\] by (metis Int_iff span_superset subsetCE) } then have "z \ f ` (rel_interior S)" using mem_rel_interior_cball[of x S] \e > 0\ x by auto } moreover { fix x assume x: "x \ rel_interior S" then obtain e2 where e2: "e2 > 0" "cball x e2 \ affine hull S \ S" using rel_interior_cball[of S] by auto have "x \ S" using x rel_interior_subset by auto then have *: "f x \ f ` S" by auto have "\x\span S. f x = 0 \ x = 0" using assms subspace_span linear_conv_bounded_linear[of f] linear_injective_on_subspace_0[of f "span S"] by auto then obtain e1 where e1: "e1 > 0" "\x \ span S. e1 * norm x \ norm (f x)" using assms injective_imp_isometric[of "span S" f] subspace_span[of S] closed_subspace[of "span S"] by auto define e where "e = e1 * e2" hence "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball (f x) e \ affine hull (f ` S)" then have "y \ f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto then obtain xy where xy: "xy \ affine hull S" "f xy = y" by auto with y have "norm (f x - f xy) \ e1 * e2" using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto moreover have "f x - f xy = f (x - xy)" using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto moreover have *: "x - xy \ span S" using subspace_diff[of "span S" x xy] subspace_span \x \ S\ xy affine_hull_subset_span[of S] span_superset by auto moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))" using e1 by auto ultimately have "e1 * norm (x - xy) \ e1 * e2" by auto then have "xy \ cball x e2" using cball_def[of x e2] dist_norm[of x xy] e1 by auto then have "y \ f ` S" using xy e2 by auto } then have "f x \ rel_interior (f ` S)" using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \e > 0\ by auto } ultimately show ?thesis by auto qed lemma rel_interior_injective_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "bounded_linear f" and "inj f" shows "rel_interior (f ` S) = f ` (rel_interior S)" using assms rel_interior_injective_on_span_linear_image[of f S] subset_inj_on[of f "UNIV" "span S"] by auto subsection\<^marker>\tag unimportant\ \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (convex hull S)" proof (clarsimp simp: open_contains_cball convex_hull_explicit) fix T and u :: "'a\real" assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1" from assms[unfolded open_contains_cball] obtain b where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis have "b ` T \ {}" using obt by auto define i where "i = b ` T" let ?\ = "\y. \F. finite F \ F \ S \ (\u. (\x\F. 0 \ u x) \ sum u F = 1 \ (\v\F. u v *\<^sub>R v) = y)" let ?a = "\v\T. u v *\<^sub>R v" show "\e > 0. cball ?a e \ {y. ?\ y}" proof (intro exI subsetI conjI) show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` T\{}\] using b \T\S\ by auto next fix y assume "y \ cball ?a (Min i)" then have y: "norm (?a - y) \ Min i" unfolding dist_norm[symmetric] by auto { fix x assume "x \ T" then have "Min i \ b x" by (simp add: i_def obt(1)) then have "x + (y - ?a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto moreover have "x \ S" using \x\T\ \T\S\ by auto ultimately have "x + (y - ?a) \ S" using y b by blast } moreover have *: "inj_on (\v. v + (y - ?a)) T" unfolding inj_on_def by auto have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" unfolding sum.reindex[OF *] o_def using obt(4) by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) ultimately show "y \ {y. ?\ y}" proof (intro CollectI exI conjI) show "finite ((\v. v + (y - ?a)) ` T)" by (simp add: obt(1)) show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1" unfolding sum.reindex[OF *] o_def using obt(4) by auto qed (use obt(1, 3) in auto) qed qed lemma compact_convex_combinations: fixes S T :: "'a::real_normed_vector set" assumes "compact S" "compact T" shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}" proof - let ?X = "{0..1} \ S \ T" let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X" by force have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) with assms show ?thesis by (simp add: * compact_Times compact_continuous_image) qed lemma finite_imp_compact_convex_hull: fixes S :: "'a::real_normed_vector set" assumes "finite S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms show ?thesis proof (induct rule: finite_ne_induct) case (singleton x) show ?case by simp next case (insert x A) let ?f = "\(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" let ?T = "{0..1::real} \ (convex hull A)" have "continuous_on ?T ?f" unfolding split_def continuous_on by (intro ballI tendsto_intros) moreover have "compact ?T" by (intro compact_Times compact_Icc insert) ultimately have "compact (?f ` ?T)" by (rule compact_continuous_image) also have "?f ` ?T = convex hull (insert x A)" unfolding convex_hull_insert [OF \A \ {}\] apply safe apply (rule_tac x=a in exI, simp) apply (rule_tac x="1 - a" in exI, simp, fast) apply (rule_tac x="(u, b)" in image_eqI, simp_all) done finally show "compact (convex hull (insert x A))" . qed qed lemma compact_convex_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis using compact_empty by simp next case False then obtain w where "w \ S" by auto show ?thesis unfolding caratheodory[of S] proof (induct ("DIM('a) + 1")) case 0 have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}" using compact_empty by auto from 0 show ?case unfolding * by simp next case (Suc n) show ?case proof (cases "n = 0") case True have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "x \ S" proof (cases "card T = 0") case True then show ?thesis using T(4) unfolding card_0_eq[OF T(1)] by simp next case False then have "card T = Suc 0" using T(3) \n=0\ by auto then obtain a where "T = {a}" unfolding card_Suc_eq by auto then show ?thesis using T(2,4) by simp qed next fix x assume "x\S" then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="{x}" in exI) unfolding convex_hull_singleton apply auto done qed then show ?thesis using assms by simp next case False have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ {x. \T. finite T \ T \ S \ card T \ n \ x \ convex hull T}}" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T" by auto moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T" apply (rule convexD_alt) using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex] using obt(7) and hull_mono[of T "insert u T"] apply auto done ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="insert u T" in exI) apply (auto simp: card_insert_if) done next fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" proof (cases "card T = Suc n") case False then have "card T \ n" using T(3) by auto then show ?thesis apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using \w\S\ and T apply (auto intro!: exI[where x=T]) done next case True then obtain a u where au: "T = insert a u" "a\u" apply (drule_tac card_eq_SucD, auto) done show ?thesis proof (cases "u = {}") case True then have "x = a" using T(4)[unfolded au] by auto show ?thesis unfolding \x = a\ apply (rule_tac x=a in exI) apply (rule_tac x=a in exI) apply (rule_tac x=1 in exI) using T and \n \ 0\ unfolding au apply (auto intro!: exI[where x="{a}"]) done next case False obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using T(4)[unfolded au convex_hull_insert[OF False]] by auto have *: "1 - vx = ux" using obt(3) by auto show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=b in exI) apply (rule_tac x=vx in exI) using obt and T(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] apply (auto intro!: exI[where x=u]) done qed qed qed then show ?thesis using compact_convex_combinations[OF assms Suc] by simp qed qed qed subsection\<^marker>\tag unimportant\ \Extremal points of a simplex are some vertices\ lemma dist_increases_online: fixes a b d :: "'a::real_inner" assumes "d \ 0" shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" proof (cases "inner a d - inner b d > 0") case True then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" apply (rule_tac add_pos_pos) using assms apply auto done then show ?thesis apply (rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done next case False then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" apply (rule_tac add_pos_nonneg) using assms apply auto done then show ?thesis apply (rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done qed lemma norm_increases_online: fixes d :: "'a::real_inner" shows "d \ 0 \ norm (a + d) > norm a \ norm(a - d) > norm a" using dist_increases_online[of d a 0] unfolding dist_norm by auto lemma simplex_furthest_lt: fixes S :: "'a::real_inner set" assumes "finite S" shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))" using assms proof induct fix x S assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))" show "\xa\convex hull insert x S. xa \ insert x S \ (\y\convex hull insert x S. norm (xa - a) < norm (y - a))" proof (intro impI ballI, cases "S = {}") case False fix y assume y: "y \ convex hull insert x S" "y \ insert x S" obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto show "\z\convex hull insert x S. norm (y - a) < norm (z - a)" proof (cases "y \ convex hull S") case True then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)" using as(3)[THEN bspec[where x=y]] and y(2) by auto then show ?thesis apply (rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] apply auto done next case False show ?thesis using obt(3) proof (cases "u = 0", case_tac[!] "v = 0") assume "u = 0" "v \ 0" then have "y = b" using obt by auto then show ?thesis using False and obt(4) by auto next assume "u \ 0" "v = 0" then have "y = x" using obt by auto then show ?thesis using y(2) by auto next assume "u \ 0" "v \ 0" then obtain w where w: "w>0" "w b" proof assume "x = b" then have "y = b" unfolding obt(5) using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) then show False using obt(4) and False by simp qed then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto show ?thesis using dist_increases_online[OF *, of a y] proof (elim disjE) assume "dist a y < dist a (y + w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u + w \ 0" "v - w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto next assume "dist a y < dist a (y - w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u - w \ 0" "v + w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto qed qed auto qed qed auto qed (auto simp: assms) lemma simplex_furthest_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)" using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \finite S\], of a] unfolding dist_commute[of a] unfolding dist_norm by auto show ?thesis proof (cases "x \ S") case False then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto then show ?thesis using x(2)[THEN bspec[where x=y]] by auto next case True with x show ?thesis by auto qed qed lemma simplex_furthest_le_exists: fixes S :: "('a::real_inner) set" shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)" using simplex_furthest_le[of S] by (cases "S = {}") auto lemma simplex_extremal_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain u v where obt: "u \ convex hull S" "v \ convex hull S" "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)" using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) then show ?thesis proof (cases "u\S \ v\S", elim disjE) assume "u \ S" then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto then show ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto next assume "v \ S" then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto then show ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) by (auto simp: norm_minus_commute) qed auto qed lemma simplex_extremal_le_exists: fixes S :: "'a::real_inner set" shows "finite S \ x \ convex hull S \ y \ convex hull S \ \u\S. \v\S. norm (x - y) \ norm (u - v)" using convex_hull_empty simplex_extremal_le[of S] by(cases "S = {}") auto subsection \Closest point of a convex set is unique, with a continuous projection\ definition\<^marker>\tag important\ closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: assumes "closed S" and "S \ {}" shows "closest_point S a \ S" and "\y\S. dist a (closest_point S a) \ dist a y" unfolding closest_point_def apply(rule_tac[!] someI2_ex) apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) done lemma closest_point_in_set: "closed S \ S \ {} \ closest_point S a \ S" by (meson closest_point_exists) lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x" using closest_point_exists[of S] by auto lemma closest_point_self: assumes "x \ S" shows "closest_point S x = x" unfolding closest_point_def apply (rule some1_equality, rule ex1I[of _ x]) using assms apply auto done lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S" using closest_point_in_set[of S x] closest_point_self[of x S] by auto lemma closer_points_lemma: assumes "inner y z > 0" shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y" proof - have z: "inner z z > 0" unfolding inner_gt_zero_iff using assms by auto have "norm (v *\<^sub>R z - y) < norm y" if "0 < v" and "v \ inner y z / inner z z" for v unfolding norm_lt using z assms that by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) then show ?thesis using assms z by (rule_tac x = "inner y z / inner z z" in exI) auto qed lemma closer_point_lemma: assumes "inner (y - x) (z - x) > 0" shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y" proof - obtain u where "u > 0" and u: "\v>0. v \ u \ norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" using closer_points_lemma[OF assms] by auto show ?thesis apply (rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and \u > 0\ unfolding dist_norm by (auto simp: norm_minus_commute field_simps) qed lemma any_closest_point_dot: assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" shows "inner (a - x) (y - x) \ 0" proof (rule ccontr) assume "\ ?thesis" then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \ S" using convexD_alt[OF assms(1,3,4), of u] using u by auto then show False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp: dist_commute algebra_simps) qed lemma any_closest_point_unique: fixes x :: "'a::real_inner" assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] unfolding norm_pths(1) and norm_le_square by (auto simp: algebra_simps) lemma closest_point_unique: assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z" shows "x = closest_point S a" using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] using closest_point_exists[OF assms(2)] and assms(3) by auto lemma closest_point_dot: assumes "convex S" "closed S" "x \ S" shows "inner (a - closest_point S a) (x - closest_point S a) \ 0" apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) using closest_point_exists[OF assms(2)] and assms(3) apply auto done lemma closest_point_lt: assumes "convex S" "closed S" "x \ S" "x \ closest_point S a" shows "dist a (closest_point S a) < dist a x" apply (rule ccontr) apply (rule_tac notE[OF assms(4)]) apply (rule closest_point_unique[OF assms(1-3), of a]) using closest_point_le[OF assms(2), of _ a] apply fastforce done lemma setdist_closest_point: "\closed S; S \ {}\ \ setdist {a} S = dist a (closest_point S a)" apply (rule setdist_unique) using closest_point_le apply (auto simp: closest_point_in_set) done lemma closest_point_lipschitz: assumes "convex S" and "closed S" "S \ {}" shows "dist (closest_point S x) (closest_point S y) \ dist x y" proof - have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0" and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0" apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) using closest_point_exists[OF assms(2-3)] apply auto done then show ?thesis unfolding dist_norm and norm_le using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous (at x) (closest_point S)" unfolding continuous_at_eps_delta using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto lemma continuous_on_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous_on t (closest_point S)" by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) proposition closest_point_in_rel_interior: assumes "closed S" "S \ {}" and x: "x \ affine hull S" shows "closest_point S x \ rel_interior S \ x \ rel_interior S" proof (cases "x \ S") case True then show ?thesis by (simp add: closest_point_self) next case False then have "False" if asm: "closest_point S x \ rel_interior S" proof - obtain e where "e > 0" and clox: "closest_point S x \ S" and e: "cball (closest_point S x) e \ affine hull S \ S" using asm mem_rel_interior_cball by blast then have clo_notx: "closest_point S x \ x" using \x \ S\ by auto define y where "y \ closest_point S x - (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" by (simp add: y_def algebra_simps) then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" by simp also have "\ < norm(x - closest_point S x)" using clo_notx \e > 0\ by (auto simp: mult_less_cancel_right2 field_split_simps) finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . have "y \ affine hull S" unfolding y_def by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) moreover have "dist (closest_point S x) y \ e" using \e > 0\ by (auto simp: y_def min_mult_distrib_right) ultimately have "y \ S" using subsetD [OF e] by simp then have "dist x (closest_point S x) \ dist x y" by (simp add: closest_point_le \closed S\) with no_less show False by (simp add: dist_norm) qed moreover have "x \ rel_interior S" using rel_interior_subset False by blast ultimately show ?thesis by blast qed subsubsection\<^marker>\tag unimportant\ \Various point-to-set separating/supporting hyperplane theorems\ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "S \ {}" and "z \ S" shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)" proof - obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2-3)]) show ?thesis proof (intro exI bexI conjI ballI) show "(y - z) \ z < (y - z) \ y" by (metis \y \ S\ assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x proof (rule ccontr) have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" using assms(1)[unfolded convex_alt] and y and \x\S\ and \y\S\ by auto assume "\ (y - z) \ y \ (y - z) \ x" then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] by (auto simp: inner_diff) then show False using *[of v] by (auto simp: dist_commute algebra_simps) qed qed (use \y \ S\ in auto) qed lemma separating_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "z \ S" shows "\a b. inner a z < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True then show ?thesis by (simp add: gt_ex) next case False obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2) False]) show ?thesis proof (intro exI conjI ballI) show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" using \y\S\ \z\S\ by auto next fix x assume "x \ S" have "False" if *: "0 < inner (z - y) (x - y)" proof - obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" using * closer_point_lemma by blast then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\] using \x\S\ \y\S\ by (auto simp: dist_commute algebra_simps) qed moreover have "0 < (norm (y - z))\<^sup>2" using \y\S\ \z\S\ by auto then have "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) qed qed lemma separating_hyperplane_closed_0: assumes "convex (S::('a::euclidean_space) set)" and "closed S" and "0 \ S" shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True have "(SOME i. i\Basis) \ (0::'a)" by (metis Basis_zero SOME_Basis) then show ?thesis using True zero_less_one by blast next case False then show ?thesis using False using separating_hyperplane_closed_point[OF assms] by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) qed subsubsection\<^marker>\tag unimportant\ \Now set-to-set for closed/compact sets\ lemma separating_hyperplane_closed_compact: fixes S :: "'a::euclidean_space set" assumes "convex S" and "closed S" and "convex T" and "compact T" and "T \ {}" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof (cases "S = {}") case True obtain b where b: "b > 0" "\x\T. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto obtain z :: 'a where z: "norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto then have "z \ T" using b(2)[THEN bspec[where x=z]] by auto then obtain a b where ab: "inner a z < b" "\x\T. b < inner a x" using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto then show ?thesis using True by auto next case False then obtain y where "y \ S" by auto obtain a b where "0 < b" "\x \ (\x\ S. \y \ T. {x - y}). b < inner a x" using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] using closed_compact_differences[OF assms(2,4)] using assms(6) by auto then have ab: "\x\S. \y\T. b + inner a y < inner a x" apply - apply rule apply rule apply (erule_tac x="x - y" in ballE) apply (auto simp: inner_diff) done define k where "k = (SUP x\T. a \ x)" show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-(k + b / 2)" in exI) apply (intro conjI ballI) unfolding inner_minus_left and neg_less_iff_less proof - fix x assume "x \ T" then have "inner a x - b / 2 < k" unfolding k_def proof (subst less_cSUP_iff) show "T \ {}" by fact show "bdd_above ((\) a ` T)" using ab[rule_format, of y] \y \ S\ by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) qed (auto intro!: bexI[of _ x] \0) then show "inner a x < k + b / 2" by auto next fix x assume "x \ S" then have "k \ inner a x - b" unfolding k_def apply (rule_tac cSUP_least) using assms(5) using ab[THEN bspec[where x=x]] apply auto done then show "k + b / 2 < inner a x" using \0 < b\ by auto qed qed lemma separating_hyperplane_compact_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "S \ {}" and "convex T" and "closed T" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof - obtain a b where "(\x\T. inner a x < b) \ (\x\S. b < inner a x)" using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto then show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-b" in exI, auto) done qed subsubsection\<^marker>\tag unimportant\ \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: assumes "convex S" "(0::'a::euclidean_space) \ S" shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)" proof - let ?k = "\c. {x::'a. 0 \ inner c x}" have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f proof - obtain c where c: "f = ?k ` c" "c \ S" "finite c" using finite_subset_image[OF as(2,1)] by auto then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) using subset_hull[of convex, OF assms(1), symmetric, of c] by force then have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)" apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) using hull_subset[of c convex] unfolding subset_eq and inner_scaleR by (auto simp: inner_commute del: ballE elim!: ballE) then show "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball sphere_def dist_norm by auto qed have "frontier (cball 0 1) \ (\(?k ` S)) \ {}" apply (rule compact_imp_fip) apply (rule compact_frontier[OF compact_cball]) using * closed_halfspace_ge by auto then obtain x where "norm x = 1" "\y\S. x\?k y" unfolding frontier_cball dist_norm sphere_def by auto then show ?thesis by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) qed lemma separating_hyperplane_sets: fixes S T :: "'a::euclidean_space set" assumes "convex S" and "convex T" and "S \ {}" and "T \ {}" and "S \ T = {}" shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)" proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x" using assms(3-5) by force then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x" by (force simp: inner_diff) then have bdd: "bdd_above (((\) a)`S)" using \T \ {}\ by (auto intro: bdd_aboveI2[OF *]) show ?thesis using \a\0\ by (intro exI[of _ a] exI[of _ "SUP x\S. a \ x"]) (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \S \ {}\ *) qed subsection\<^marker>\tag unimportant\ \More convexity generalities\ lemma convex_closure [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (closure S)" apply (rule convexI) apply (unfold closure_sequential, elim exE) apply (rule_tac x="\n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) apply (rule,rule) apply (rule convexD [OF assms]) apply (auto del: tendsto_const intro!: tendsto_intros) done lemma convex_interior [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (interior S)" unfolding convex_alt Ball_def mem_interior proof clarify fix x y u assume u: "0 \ u" "u \ (1::real)" fix e d assume ed: "ball x e \ S" "ball y d \ S" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S" proof (intro exI conjI subsetI) fix z assume "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S" apply (rule_tac assms[unfolded convex_alt, rule_format]) using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm apply (auto simp: algebra_simps) done then show "z \ S" using u by (auto simp: algebra_simps) qed(insert u ed(3-4), auto) qed lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}" using hull_subset[of S convex] convex_hull_empty by auto subsection\<^marker>\tag unimportant\ \Convex set as intersection of halfspaces\ lemma convex_halfspace_intersection: fixes s :: "('a::euclidean_space) set" assumes "closed s" "convex s" shows "s = \{h. s \ h \ (\a b. h = {x. inner a x \ b})}" apply (rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply (rule,rule,erule conjE) proof - fix x assume "\xa. s \ xa \ (\a b. xa = {x. inner a x \ b}) \ x \ xa" then have "\a b. s \ {x. inner a x \ b} \ x \ {x. inner a x \ b}" by blast then show "x \ s" apply (rule_tac ccontr) apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) apply (erule exE)+ apply (erule_tac x="-a" in allE) apply (erule_tac x="-b" in allE, auto) done qed auto subsection\<^marker>\tag unimportant\ \Convexity of general and special intervals\ lemma is_interval_convex: fixes S :: "'a::euclidean_space set" assumes "is_interval S" shows "convex S" proof (rule convexI) fix x y and u v :: real assume as: "x \ S" "y \ S" "0 \ u" "0 \ v" "u + v = 1" then have *: "u = 1 - v" "1 - v \ 0" and **: "v = 1 - u" "1 - u \ 0" by auto { fix a b assume "\ b \ u * a + v * b" then have "u * a < (1 - v) * b" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) *(2) apply (rule_tac mult_left_less_imp_less[of "1 - v"]) apply (auto simp: field_simps) done then have "a \ u * a + v * b" unfolding * using as(4) by (auto simp: field_simps intro!:mult_right_mono) } moreover { fix a b assume "\ u * a + v * b \ a" then have "v * b > (1 - u) * a" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) apply (rule_tac mult_left_less_imp_less) apply (auto simp: field_simps) done then have "u * a + v * b \ b" unfolding ** using **(2) as(3) by (auto simp: field_simps intro!:mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ S" apply - apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) DIM_positive[where 'a='a] apply (auto simp: inner_simps) done qed lemma is_interval_connected: fixes S :: "'a::euclidean_space set" shows "is_interval S \ connected S" using is_interval_convex convex_connected by auto lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" apply (rule_tac[!] is_interval_convex)+ using is_interval_box is_interval_cbox apply auto done text\A non-singleton connected set is perfect (i.e. has no isolated points). \ lemma connected_imp_perfect: fixes a :: "'a::metric_space" assumes "connected S" "a \ S" and S: "\x. S \ {x}" shows "a islimpt S" proof - have False if "a \ T" "open T" "\y. \y \ S; y \ T\ \ y = a" for T proof - obtain e where "e > 0" and e: "cball a e \ T" using \open T\ \a \ T\ by (auto simp: open_contains_cball) have "openin (top_of_set S) {a}" unfolding openin_open using that \a \ S\ by blast moreover have "closedin (top_of_set S) {a}" by (simp add: assms) ultimately show "False" using \connected S\ connected_clopen S by blast qed then show ?thesis unfolding islimpt_def by blast qed lemma connected_imp_perfect_aff_dim: "\connected S; aff_dim S \ 0; a \ S\ \ a islimpt S" using aff_dim_sing connected_imp_perfect by blast subsection\<^marker>\tag unimportant\ \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ lemma mem_is_interval_1_I: fixes a b c::real assumes "is_interval S" assumes "a \ S" "c \ S" assumes "a \ b" "b \ c" shows "b \ S" using assms is_interval_1 by blast lemma is_interval_connected_1: fixes s :: "real set" shows "is_interval s \ connected s" apply rule apply (rule is_interval_connected, assumption) unfolding is_interval_1 apply rule apply rule apply rule apply rule apply (erule conjE) apply (rule ccontr) proof - fix a b x assume as: "connected s" "a \ s" "b \ s" "a \ x" "x \ b" "x \ s" then have *: "a < x" "x < b" unfolding not_le [symmetric] by auto let ?halfl = "{.. s" with \x \ s\ have "x \ y" by auto then have "y \ ?halfr \ ?halfl" by auto } moreover have "a \ ?halfl" "b \ ?halfr" using * by auto then have "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto ultimately show False apply (rule_tac notE[OF as(1)[unfolded connected_def]]) apply (rule_tac x = ?halfl in exI) apply (rule_tac x = ?halfr in exI, rule) apply (rule open_lessThan, rule) apply (rule open_greaterThan, auto) done qed lemma is_interval_convex_1: fixes s :: "real set" shows "is_interval s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real by (metis connected_ball is_interval_connected_1) lemma connected_compact_interval_1: "connected S \ compact S \ (\a b. S = {a..b::real})" by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) lemma connected_convex_1: fixes s :: "real set" shows "connected s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_convex_1_gen: fixes s :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected s \ convex s" proof - obtain f:: "'a \ real" where linf: "linear f" and "inj f" using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, where 'a='a and 'b=real] unfolding Euclidean_Space.dim_UNIV by (auto simp: assms) then have "f -` (f ` s) = s" by (simp add: inj_vimage_image_eq) then show ?thesis by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) qed lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real by (simp add: is_interval_convex_1) lemma [simp]: fixes r s::real shows is_interval_io: "is_interval {..\tag unimportant\ \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f a)\k \ y" "y \ (f b)\k" shows "\x\{a..b}. (f x)\k = y" proof - have "f a \ f ` cbox a b" "f b \ f ` cbox a b" apply (rule_tac[!] imageI) using assms(1) apply auto done then show ?thesis using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] by (simp add: connected_continuous_image assms) qed lemma ivt_increasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f b)\k \ y" and "y \ (f a)\k" shows "\x\{a..b}. (f x)\k = y" apply (subst neg_equal_iff_equal[symmetric]) using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] using assms using continuous_on_minus apply auto done lemma ivt_decreasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) subsection\<^marker>\tag unimportant\ \A bound within an interval\ lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" shows "convex hull {x, y} = cbox x y" proof (rule hull_unique) show "{x, y} \ cbox x y" using \x \ y\ by auto show "convex (cbox x y)" by (rule convex_box) next fix S assume "{x, y} \ S" and "convex S" then show "cbox x y \ S" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed lemma unit_interval_convex_hull: "cbox (0::'a::euclidean_space) One = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") proof - have One[simp]: "\i. i \ Basis \ One \ i = 1" by (simp add: inner_sum_left sum.If_cases inner_Basis) have "?int = {x. \i\Basis. x \ i \ cbox 0 1}" by (auto simp: cbox_def) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` cbox 0 1)" by (simp only: box_eq_set_sum_Basis) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" by (simp add: convex_hull_linear_image) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" by (simp only: convex_hull_set_sum) also have "\ = convex hull {x. \i\Basis. x\i \ {0, 1}}" by (simp only: box_eq_set_sum_Basis) also have "convex hull {x. \i\Basis. x\i \ {0, 1}} = convex hull ?points" by simp finally show ?thesis . qed text \And this is a finite set of vertices.\ lemma unit_cube_convex_hull: obtains S :: "'a::euclidean_space set" where "finite S" and "cbox 0 (\Basis) = convex hull S" proof show "finite {x::'a. \i\Basis. x \ i = 0 \ x \ i = 1}" proof (rule finite_subset, clarify) show "finite ((\S. \i\Basis. (if i \ S then 1 else 0) *\<^sub>R i) ` Pow Basis)" using finite_Basis by blast fix x :: 'a assume as: "\i\Basis. x \ i = 0 \ x \ i = 1" show "x \ (\S. \i\Basis. (if i\S then 1 else 0) *\<^sub>R i) ` Pow Basis" apply (rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) using as apply (subst euclidean_eq_iff, auto) done qed show "cbox 0 One = convex hull {x. \i\Basis. x \ i = 0 \ x \ i = 1}" using unit_interval_convex_hull by blast qed text \Hence any cube (could do any nonempty interval).\ lemma cube_convex_hull: assumes "d > 0" obtains S :: "'a::euclidean_space set" where "finite S" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull S" proof - let ?d = "(\i\Basis. d *\<^sub>R i)::'a" have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" proof (intro set_eqI iffI) fix y assume "y \ cbox (x - ?d) (x + ?d)" then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" using assms by (simp add: mem_box inner_simps) (simp add: field_simps) with \0 < d\ show "y \ (\y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) next fix y assume "y \ (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" then obtain z where z: "z \ cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" by auto then show "y \ cbox (x - ?d) (x + ?d)" using z assms by (auto simp: mem_box inner_simps) qed obtain S where "finite S" "cbox 0 (\Basis::'a) = convex hull S" using unit_cube_convex_hull by auto then show ?thesis by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed subsection\<^marker>\tag unimportant\\Representation of any interval as a finite convex hull\ lemma image_stretch_interval: "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = (if (cbox a b) = {} then {} else cbox (\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k::'a) (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k))" proof cases assume *: "cbox a b \ {}" show ?thesis unfolding box_ne_empty if_not_P[OF *] apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) apply (subst choice_Basis_iff[symmetric]) proof (intro allI ball_cong refl) fix x i :: 'a assume "i \ Basis" with * have a_le_b: "a \ i \ b \ i" unfolding box_ne_empty by auto show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" proof (cases "m i = 0") case True with a_le_b show ?thesis by auto next case False then have *: "\a b. a = m i * b \ b = a / m i" by (auto simp: field_simps) from False have "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) with False show ?thesis using a_le_b unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps) qed qed qed simp lemma interval_image_stretch_interval: "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" unfolding image_stretch_interval by auto lemma cbox_translation: "cbox (c + a) (c + b) = image (\x. c + x) (cbox a b)" using image_affinity_cbox [of 1 c a b] using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] by (auto simp: inner_left_distrib add.commute) lemma cbox_image_unit_interval: fixes a :: "'a::euclidean_space" assumes "cbox a b \ {}" shows "cbox a b = (+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` cbox 0 One" using assms apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) done lemma closed_interval_as_convex_hull: fixes a :: "'a::euclidean_space" obtains S where "finite S" "cbox a b = convex hull S" proof (cases "cbox a b = {}") case True with convex_hull_empty that show ?thesis by blast next case False obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" by (blast intro: unit_cube_convex_hull) have lin: "linear (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k)" by (rule linear_compose_sum) (auto simp: algebra_simps linearI) have "finite ((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` S)" by (rule finite_imageI \finite S\)+ then show ?thesis apply (rule that) apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) done qed subsection\<^marker>\tag unimportant\ \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: fixes S :: "('a::real_normed_vector) set" assumes "open S" and "convex_on S f" and "\x\S. \f x\ \ b" shows "continuous_on S f" apply (rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof (rule,rule,rule) fix x and e :: real assume "x \ S" "e > 0" define B where "B = \b\ + 1" then have B: "0 < B""\x. x\S \ \f x\ \ B" using assms(3) by auto obtain k where "k > 0" and k: "cball x k \ S" using \x \ S\ assms(1) open_contains_cball_eq by blast show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" proof (intro exI conjI allI impI) fix y assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" show "\f y - f x\ < e" proof (cases "y = x") case False define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ by (auto simp:field_simps) have "y \ S" apply (rule k[THEN subsetD]) unfolding mem_cball dist_norm apply (rule order_trans[of _ "2 * norm (x - y)"]) using as by (auto simp: field_simps norm_minus_commute) { define w where "w = x + t *\<^sub>R (y - x)" have "w \ S" using \k>0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = 0" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have 2: "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) have "f y - f x \ (f w - f x) / t" using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] using \0 < t\ \2 < t\ and \x \ S\ \w \ S\ by (auto simp:field_simps) also have "... < e" using B(2)[OF \w\S\] and B(2)[OF \x\S\] 2 \t > 0\ by (auto simp: field_simps) finally have th1: "f y - f x < e" . } moreover { define w where "w = x - t *\<^sub>R (y - x)" have "w \ S" using \k > 0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = x" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) then have *: "(f w - f y) / t < e" using B(2)[OF \w\S\] and B(2)[OF \y\S\] using \t > 0\ by (auto simp:field_simps) have "f x \ 1 / (1 + t) * f w + (t / (1 + t)) * f y" using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] using \0 < t\ \2 < t\ and \y \ S\ \w \ S\ by (auto simp:field_simps) also have "\ = (f w + t * f y) / (1 + t)" using \t > 0\ by (simp add: add_divide_distrib) also have "\ < e + f y" using \t > 0\ * \e > 0\ by (auto simp: field_simps) finally have "f x - f y < e" by auto } ultimately show ?thesis by auto qed (insert \0, auto) qed (insert \0 \0 \0, auto simp: field_simps) qed subsection\<^marker>\tag unimportant\ \Upper bound on a ball implies upper and lower bounds\ lemma convex_bounds_lemma: fixes x :: "'a::real_normed_vector" assumes "convex_on (cball x e) f" and "\y \ cball x e. f y \ b" shows "\y \ cball x e. \f y\ \ b + 2 * \f x\" apply rule proof (cases "0 \ e") case True fix y assume y: "y \ cball x e" define z where "z = 2 *\<^sub>R x - y" have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp: algebra_simps) then show "\f y\ \ b + 2 * \f x\" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by (auto simp:field_simps) next case False fix y assume "y \ cball x e" then have "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) then show "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed subsubsection\<^marker>\tag unimportant\ \Hence a convex function on an open set is continuous\ lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" by auto lemma convex_on_continuous: assumes "open (s::('a::euclidean_space) set)" "convex_on s f" shows "continuous_on s f" unfolding continuous_on_eq_continuous_at[OF assms(1)] proof note dimge1 = DIM_positive[where 'a='a] fix x assume "x \ s" then obtain e where e: "cball x e \ s" "e > 0" using assms(1) unfolding open_contains_cball by auto define d where "d = e / real DIM('a)" have "0 < d" unfolding d_def using \e > 0\ dimge1 by auto let ?d = "(\i\Basis. d *\<^sub>R i)::'a" obtain c where c: "finite c" and c1: "convex hull c \ cball x e" and c2: "cball x d \ convex hull c" proof define c where "c = (\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d})" show "finite c" unfolding c_def by (simp add: finite_set_sum) have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" unfolding box_eq_set_sum_Basis unfolding c_def convex_hull_set_sum apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_add_left) apply (rule sum.cong [OF refl]) apply (rule image_cong [OF _ refl]) apply (rule convex_hull_eq_real_cbox) apply (cut_tac \0 < d\, simp) done then have 2: "convex hull c = {a. \i\Basis. a \ i \ cball (x \ i) d}" by (simp add: dist_norm abs_le_iff algebra_simps) show "cball x d \ convex hull c" unfolding 2 by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) have e': "e = (\(i::'a)\Basis. d)" by (simp add: d_def DIM_positive) show "convex hull c \ cball x e" unfolding 2 apply clarsimp apply (subst euclidean_dist_l2) apply (rule order_trans [OF L2_set_le_sum]) apply (rule zero_le_dist) unfolding e' apply (rule sum_mono, simp) done qed define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" apply(rule convex_on_subset[OF assms(2)]) apply(rule subset_trans[OF c1 e(1)]) done then have k: "\y\convex hull c. f y \ k" apply (rule_tac convex_on_convex_hull_bound, assumption) by (simp add: k_def c) have "e \ e * real DIM('a)" using e(2) real_of_nat_ge_one_iff by auto then have "d \ e" by (simp add: d_def field_split_simps) then have dsube: "cball x d \ cball x e" by (rule subset_cball) have conv: "convex_on (cball x d) f" using \convex_on (convex hull c) f\ c2 convex_on_subset by blast then have "\y\cball x d. \f y\ \ k + 2 * \f x\" by (rule convex_bounds_lemma) (use c2 k in blast) then have "continuous_on (ball x d) f" apply (rule_tac convex_on_bounded_continuous) apply (rule open_ball, rule convex_on_subset[OF conv]) apply (rule ball_subset_cball, force) done then show "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using \d > 0\ by auto qed + +section \Line Segments\ + +subsection \Midpoint\ + +definition\<^marker>\tag important\ midpoint :: "'a::real_vector \ 'a \ 'a" + where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" + +lemma midpoint_idem [simp]: "midpoint x x = x" + unfolding midpoint_def by simp + +lemma midpoint_sym: "midpoint a b = midpoint b a" + unfolding midpoint_def by (auto simp add: scaleR_right_distrib) + +lemma midpoint_eq_iff: "midpoint a b = c \ a + b = c + c" +proof - + have "midpoint a b = c \ scaleR 2 (midpoint a b) = scaleR 2 c" + by simp + then show ?thesis + unfolding midpoint_def scaleR_2 [symmetric] by simp +qed + +lemma + fixes a::real + assumes "a \ b" shows ge_midpoint_1: "a \ midpoint a b" + and le_midpoint_1: "midpoint a b \ b" + by (simp_all add: midpoint_def assms) + +lemma dist_midpoint: + fixes a b :: "'a::real_normed_vector" shows + "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) + "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) + "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) + "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) +proof - + have *: "\x y::'a. 2 *\<^sub>R x = - y \ norm x = (norm y) / 2" + unfolding equation_minus_iff by auto + have **: "\x y::'a. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" + by auto + note scaleR_right_distrib [simp] + show ?t1 + unfolding midpoint_def dist_norm + apply (rule **) + apply (simp add: scaleR_right_diff_distrib) + apply (simp add: scaleR_2) + done + show ?t2 + unfolding midpoint_def dist_norm + apply (rule *) + apply (simp add: scaleR_right_diff_distrib) + apply (simp add: scaleR_2) + done + show ?t3 + unfolding midpoint_def dist_norm + apply (rule *) + apply (simp add: scaleR_right_diff_distrib) + apply (simp add: scaleR_2) + done + show ?t4 + unfolding midpoint_def dist_norm + apply (rule **) + apply (simp add: scaleR_right_diff_distrib) + apply (simp add: scaleR_2) + done +qed + +lemma midpoint_eq_endpoint [simp]: + "midpoint a b = a \ a = b" + "midpoint a b = b \ a = b" + unfolding midpoint_eq_iff by auto + +lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b" + using midpoint_eq_iff by metis + +lemma midpoint_linear_image: + "linear f \ midpoint(f a)(f b) = f(midpoint a b)" +by (simp add: linear_iff midpoint_def) + + +subsection \Line segments\ + +definition\<^marker>\tag important\ closed_segment :: "'a::real_vector \ 'a \ 'a set" + where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" + +definition\<^marker>\tag important\ open_segment :: "'a::real_vector \ 'a \ 'a set" where + "open_segment a b \ closed_segment a b - {a,b}" + +lemmas segment = open_segment_def closed_segment_def + +lemma in_segment: + "x \ closed_segment a b \ (\u. 0 \ u \ u \ 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" + "x \ open_segment a b \ a \ b \ (\u. 0 < u \ u < 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" + using less_eq_real_def by (auto simp: segment algebra_simps) + +lemma closed_segment_linear_image: + "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f" +proof - + interpret linear f by fact + show ?thesis + by (force simp add: in_segment add scale) +qed + +lemma open_segment_linear_image: + "\linear f; inj f\ \ open_segment (f a) (f b) = f ` (open_segment a b)" + by (force simp: open_segment_def closed_segment_linear_image inj_on_def) + +lemma closed_segment_translation: + "closed_segment (c + a) (c + b) = image (\x. c + x) (closed_segment a b)" +apply safe +apply (rule_tac x="x-c" in image_eqI) +apply (auto simp: in_segment algebra_simps) +done + +lemma open_segment_translation: + "open_segment (c + a) (c + b) = image (\x. c + x) (open_segment a b)" +by (simp add: open_segment_def closed_segment_translation translation_diff) + +lemma closed_segment_of_real: + "closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y" + apply (auto simp: image_iff in_segment scaleR_conv_of_real) + apply (rule_tac x="(1-u)*x + u*y" in bexI) + apply (auto simp: in_segment) + done + +lemma open_segment_of_real: + "open_segment (of_real x) (of_real y) = of_real ` open_segment x y" + apply (auto simp: image_iff in_segment scaleR_conv_of_real) + apply (rule_tac x="(1-u)*x + u*y" in bexI) + apply (auto simp: in_segment) + done + +lemma closed_segment_Reals: + "\x \ Reals; y \ Reals\ \ closed_segment x y = of_real ` closed_segment (Re x) (Re y)" + by (metis closed_segment_of_real of_real_Re) + +lemma open_segment_Reals: + "\x \ Reals; y \ Reals\ \ open_segment x y = of_real ` open_segment (Re x) (Re y)" + by (metis open_segment_of_real of_real_Re) + +lemma open_segment_PairD: + "(x, x') \ open_segment (a, a') (b, b') + \ (x \ open_segment a b \ a = b) \ (x' \ open_segment a' b' \ a' = b')" + by (auto simp: in_segment) + +lemma closed_segment_PairD: + "(x, x') \ closed_segment (a, a') (b, b') \ x \ closed_segment a b \ x' \ closed_segment a' b'" + by (auto simp: closed_segment_def) + +lemma closed_segment_translation_eq [simp]: + "d + x \ closed_segment (d + a) (d + b) \ x \ closed_segment a b" +proof - + have *: "\d x a b. x \ closed_segment a b \ d + x \ closed_segment (d + a) (d + b)" + apply (simp add: closed_segment_def) + apply (erule ex_forward) + apply (simp add: algebra_simps) + done + show ?thesis + using * [where d = "-d"] * + by (fastforce simp add:) +qed + +lemma open_segment_translation_eq [simp]: + "d + x \ open_segment (d + a) (d + b) \ x \ open_segment a b" + by (simp add: open_segment_def) + +lemma of_real_closed_segment [simp]: + "of_real x \ closed_segment (of_real a) (of_real b) \ x \ closed_segment a b" + apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward) + using of_real_eq_iff by fastforce + +lemma of_real_open_segment [simp]: + "of_real x \ open_segment (of_real a) (of_real b) \ x \ open_segment a b" + apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE) + using of_real_eq_iff by fastforce + +lemma convex_contains_segment: + "convex S \ (\a\S. \b\S. closed_segment a b \ S)" + unfolding convex_alt closed_segment_def by auto + +lemma closed_segment_in_Reals: + "\x \ closed_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" + by (meson subsetD convex_Reals convex_contains_segment) + +lemma open_segment_in_Reals: + "\x \ open_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" + by (metis Diff_iff closed_segment_in_Reals open_segment_def) + +lemma closed_segment_subset: "\x \ S; y \ S; convex S\ \ closed_segment x y \ S" + by (simp add: convex_contains_segment) + +lemma closed_segment_subset_convex_hull: + "\x \ convex hull S; y \ convex hull S\ \ closed_segment x y \ convex hull S" + using convex_contains_segment by blast + +lemma segment_convex_hull: + "closed_segment a b = convex hull {a,b}" +proof - + have *: "\x. {x} \ {}" by auto + show ?thesis + unfolding segment convex_hull_insert[OF *] convex_hull_singleton + by (safe; rule_tac x="1 - u" in exI; force) +qed + +lemma open_closed_segment: "u \ open_segment w z \ u \ closed_segment w z" + by (auto simp add: closed_segment_def open_segment_def) + +lemma segment_open_subset_closed: + "open_segment a b \ closed_segment a b" + by (auto simp: closed_segment_def open_segment_def) + +lemma bounded_closed_segment: + fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)" + by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded) + +lemma bounded_open_segment: + fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)" + by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed]) + +lemmas bounded_segment = bounded_closed_segment open_closed_segment + +lemma ends_in_segment [iff]: "a \ closed_segment a b" "b \ closed_segment a b" + unfolding segment_convex_hull + by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) + +lemma eventually_closed_segment: + fixes x0::"'a::real_normed_vector" + assumes "open X0" "x0 \ X0" + shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" +proof - + from openE[OF assms] + obtain e where e: "0 < e" "ball x0 e \ X0" . + then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" + by (auto simp: dist_commute eventually_at) + then show ?thesis + proof eventually_elim + case (elim x) + have "x0 \ ball x0 e" using \e > 0\ by simp + from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] + have "closed_segment x0 x \ ball x0 e" . + also note \\ \ X0\ + finally show ?case . + qed +qed + +lemma segment_furthest_le: + fixes a b x y :: "'a::euclidean_space" + assumes "x \ closed_segment a b" + shows "norm (y - x) \ norm (y - a) \ norm (y - x) \ norm (y - b)" +proof - + obtain z where "z \ {a, b}" "norm (x - y) \ norm (z - y)" + using simplex_furthest_le[of "{a, b}" y] + using assms[unfolded segment_convex_hull] + by auto + then show ?thesis + by (auto simp add:norm_minus_commute) +qed + +lemma closed_segment_commute: "closed_segment a b = closed_segment b a" +proof - + have "{a, b} = {b, a}" by auto + thus ?thesis + by (simp add: segment_convex_hull) +qed + +lemma segment_bound1: + assumes "x \ closed_segment a b" + shows "norm (x - a) \ norm (b - a)" +proof - + obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" + using assms by (auto simp add: closed_segment_def) + then show "norm (x - a) \ norm (b - a)" + apply clarify + apply (auto simp: algebra_simps) + apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le) + done +qed + +lemma segment_bound: + assumes "x \ closed_segment a b" + shows "norm (x - a) \ norm (b - a)" "norm (x - b) \ norm (b - a)" +apply (simp add: assms segment_bound1) +by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1) + +lemma open_segment_commute: "open_segment a b = open_segment b a" +proof - + have "{a, b} = {b, a}" by auto + thus ?thesis + by (simp add: closed_segment_commute open_segment_def) +qed + +lemma closed_segment_idem [simp]: "closed_segment a a = {a}" + unfolding segment by (auto simp add: algebra_simps) + +lemma open_segment_idem [simp]: "open_segment a a = {}" + by (simp add: open_segment_def) + +lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \ {a,b}" + using open_segment_def by auto + +lemma convex_contains_open_segment: + "convex s \ (\a\s. \b\s. open_segment a b \ s)" + by (simp add: convex_contains_segment closed_segment_eq_open) + +lemma closed_segment_eq_real_ivl: + fixes a b::real + shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" +proof - + have "b \ a \ closed_segment b a = {b .. a}" + and "a \ b \ closed_segment a b = {a .. b}" + by (auto simp: convex_hull_eq_real_cbox segment_convex_hull) + thus ?thesis + by (auto simp: closed_segment_commute) +qed + +lemma open_segment_eq_real_ivl: + fixes a b::real + shows "open_segment a b = (if a \ b then {a<..x. (v - u) * x + u) ` {0..1}" + by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) + +lemma dist_in_closed_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \ closed_segment a b" + shows "dist x a \ dist a b \ dist x b \ dist a b" +proof (intro conjI) + obtain u where u: "0 \ u" "u \ 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" + using assms by (force simp: in_segment algebra_simps) + have "dist x a = u * dist a b" + apply (simp add: dist_norm algebra_simps x) + by (metis \0 \ u\ abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) + also have "... \ dist a b" + by (simp add: mult_left_le_one_le u) + finally show "dist x a \ dist a b" . + have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" + by (simp add: dist_norm algebra_simps x) + also have "... = (1-u) * dist a b" + proof - + have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" + using \u \ 1\ by force + then show ?thesis + by (simp add: dist_norm real_vector.scale_right_diff_distrib) + qed + also have "... \ dist a b" + by (simp add: mult_left_le_one_le u) + finally show "dist x b \ dist a b" . +qed + +lemma dist_in_open_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \ open_segment a b" + shows "dist x a < dist a b \ dist x b < dist a b" +proof (intro conjI) + obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" + using assms by (force simp: in_segment algebra_simps) + have "dist x a = u * dist a b" + apply (simp add: dist_norm algebra_simps x) + by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \0 < u\) + also have *: "... < dist a b" + by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \u < 1\) + finally show "dist x a < dist a b" . + have ab_ne0: "dist a b \ 0" + using * by fastforce + have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" + by (simp add: dist_norm algebra_simps x) + also have "... = (1-u) * dist a b" + proof - + have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" + using \u < 1\ by force + then show ?thesis + by (simp add: dist_norm real_vector.scale_right_diff_distrib) + qed + also have "... < dist a b" + using ab_ne0 \0 < u\ by simp + finally show "dist x b < dist a b" . +qed + +lemma dist_decreases_open_segment_0: + fixes x :: "'a :: euclidean_space" + assumes "x \ open_segment 0 b" + shows "dist c x < dist c 0 \ dist c x < dist c b" +proof (rule ccontr, clarsimp simp: not_less) + obtain u where u: "0 \ b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" + using assms by (auto simp: in_segment) + have xb: "x \ b < b \ b" + using u x by auto + assume "norm c \ dist c x" + then have "c \ c \ (c - x) \ (c - x)" + by (simp add: dist_norm norm_le) + moreover have "0 < x \ b" + using u x by auto + ultimately have less: "c \ b < x \ b" + by (simp add: x algebra_simps inner_commute u) + assume "dist c b \ dist c x" + then have "(c - b) \ (c - b) \ (c - x) \ (c - x)" + by (simp add: dist_norm norm_le) + then have "(b \ b) * (1 - u*u) \ 2 * (b \ c) * (1-u)" + by (simp add: x algebra_simps inner_commute) + then have "(1+u) * (b \ b) * (1-u) \ 2 * (b \ c) * (1-u)" + by (simp add: algebra_simps) + then have "(1+u) * (b \ b) \ 2 * (b \ c)" + using \u < 1\ by auto + with xb have "c \ b \ x \ b" + by (auto simp: x algebra_simps inner_commute) + with less show False by auto +qed + +proposition dist_decreases_open_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \ open_segment a b" + shows "dist c x < dist c a \ dist c x < dist c b" +proof - + have *: "x - a \ open_segment 0 (b - a)" using assms + by (metis diff_self open_segment_translation_eq uminus_add_conv_diff) + show ?thesis + using dist_decreases_open_segment_0 [OF *, of "c-a"] assms + by (simp add: dist_norm) +qed + +corollary open_segment_furthest_le: + fixes a b x y :: "'a::euclidean_space" + assumes "x \ open_segment a b" + shows "norm (y - x) < norm (y - a) \ norm (y - x) < norm (y - b)" + by (metis assms dist_decreases_open_segment dist_norm) + +corollary dist_decreases_closed_segment: + fixes a :: "'a :: euclidean_space" + assumes "x \ closed_segment a b" + shows "dist c x \ dist c a \ dist c x \ dist c b" +apply (cases "x \ open_segment a b") + using dist_decreases_open_segment less_eq_real_def apply blast +by (metis DiffI assms empty_iff insertE open_segment_def order_refl) + +lemma convex_intermediate_ball: + fixes a :: "'a :: euclidean_space" + shows "\ball a r \ T; T \ cball a r\ \ convex T" +apply (simp add: convex_contains_open_segment, clarify) +by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment) + +lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \ closed_segment a b" + apply (clarsimp simp: midpoint_def in_segment) + apply (rule_tac x="(1 + u) / 2" in exI) + apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib) + by (metis field_sum_of_halves scaleR_left.add) + +lemma notin_segment_midpoint: + fixes a :: "'a :: euclidean_space" + shows "a \ b \ a \ closed_segment (midpoint a b) b" +by (auto simp: dist_midpoint dest!: dist_in_closed_segment) + +lemma segment_to_closest_point: + fixes S :: "'a :: euclidean_space set" + shows "\closed S; S \ {}\ \ open_segment a (closest_point S a) \ S = {}" + apply (subst disjoint_iff_not_equal) + apply (clarify dest!: dist_in_open_segment) + by (metis closest_point_le dist_commute le_less_trans less_irrefl) + +lemma segment_to_point_exists: + fixes S :: "'a :: euclidean_space set" + assumes "closed S" "S \ {}" + obtains b where "b \ S" "open_segment a b \ S = {}" + by (metis assms segment_to_closest_point closest_point_exists that) + +subsubsection\More lemmas, especially for working with the underlying formula\ + +lemma segment_eq_compose: + fixes a :: "'a :: real_vector" + shows "(\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\x. a + x) o (\u. u *\<^sub>R (b - a))" + by (simp add: o_def algebra_simps) + +lemma segment_degen_1: + fixes a :: "'a :: real_vector" + shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \ a=b \ u=1" +proof - + { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b" + then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b" + by (simp add: algebra_simps) + then have "a=b \ u=1" + by simp + } then show ?thesis + by (auto simp: algebra_simps) +qed + +lemma segment_degen_0: + fixes a :: "'a :: real_vector" + shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \ a=b \ u=0" + using segment_degen_1 [of "1-u" b a] + by (auto simp: algebra_simps) + +lemma add_scaleR_degen: + fixes a b ::"'a::real_vector" + assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \ v" + shows "a=b" + by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms) + +lemma closed_segment_image_interval: + "closed_segment a b = (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}" + by (auto simp: set_eq_iff image_iff closed_segment_def) + +lemma open_segment_image_interval: + "open_segment a b = (if a=b then {} else (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})" + by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1) + +lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval + +lemma open_segment_bound1: + assumes "x \ open_segment a b" + shows "norm (x - a) < norm (b - a)" +proof - + obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \ b" + using assms by (auto simp add: open_segment_image_interval split: if_split_asm) + then show "norm (x - a) < norm (b - a)" + apply clarify + apply (auto simp: algebra_simps) + apply (simp add: scaleR_diff_right [symmetric]) + done +qed + +lemma compact_segment [simp]: + fixes a :: "'a::real_normed_vector" + shows "compact (closed_segment a b)" + by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros) + +lemma closed_segment [simp]: + fixes a :: "'a::real_normed_vector" + shows "closed (closed_segment a b)" + by (simp add: compact_imp_closed) + +lemma closure_closed_segment [simp]: + fixes a :: "'a::real_normed_vector" + shows "closure(closed_segment a b) = closed_segment a b" + by simp + +lemma open_segment_bound: + assumes "x \ open_segment a b" + shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)" +apply (simp add: assms open_segment_bound1) +by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute) + +lemma closure_open_segment [simp]: + "closure (open_segment a b) = (if a = b then {} else closed_segment a b)" + for a :: "'a::euclidean_space" +proof (cases "a = b") + case True + then show ?thesis + by simp +next + case False + have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" + apply (rule closure_injective_linear_image [symmetric]) + apply (use False in \auto intro!: injI\) + done + then have "closure + ((\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) = + (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}" + using closure_translation [of a "((\x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"] + by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image) + then show ?thesis + by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan) +qed + +lemma closed_open_segment_iff [simp]: + fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \ a = b" + by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2)) + +lemma compact_open_segment_iff [simp]: + fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \ a = b" + by (simp add: bounded_open_segment compact_eq_bounded_closed) + +lemma convex_closed_segment [iff]: "convex (closed_segment a b)" + unfolding segment_convex_hull by(rule convex_convex_hull) + +lemma convex_open_segment [iff]: "convex (open_segment a b)" +proof - + have "convex ((\u. u *\<^sub>R (b - a)) ` {0<..<1})" + by (rule convex_linear_image) auto + then have "convex ((+) a ` (\u. u *\<^sub>R (b - a)) ` {0<..<1})" + by (rule convex_translation) + then show ?thesis + by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right) +qed + +lemmas convex_segment = convex_closed_segment convex_open_segment + +lemma connected_segment [iff]: + fixes x :: "'a :: real_normed_vector" + shows "connected (closed_segment x y)" + by (simp add: convex_connected) + +lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real + by (auto simp: is_interval_convex_1) + +lemma IVT'_closed_segment_real: + fixes f :: "real \ real" + assumes "y \ closed_segment (f a) (f b)" + assumes "continuous_on (closed_segment a b) f" + shows "\x \ closed_segment a b. f x = y" + using IVT'[of f a y b] + IVT'[of "-f" a "-y" b] + IVT'[of f b y a] + IVT'[of "-f" b "-y" a] assms + by (cases "a \ b"; cases "f b \ f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus) + end diff --git a/src/HOL/Analysis/Derivative.thy b/src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy +++ b/src/HOL/Analysis/Derivative.thy @@ -1,2900 +1,2901 @@ (* Title: HOL/Analysis/Derivative.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (translation from HOL Light); tidied by LCP *) section \Derivative\ theory Derivative imports - Starlike + Convex_Euclidean_Space + Abstract_Limits Operator_Norm Uniform_Limit Bounded_Linear_Function begin declare bounded_linear_inner_left [intro] declare has_derivative_bounded_linear[dest] subsection \Derivatives\ lemma has_derivative_add_const: "(f has_derivative f') net \ ((\x. f x + c) has_derivative f') net" by (intro derivative_eq_intros) auto subsection\<^marker>\tag unimportant\ \Derivative with composed bilinear function\ text \More explicit epsilon-delta forms.\ proposition has_derivative_within': "(f has_derivative f')(at x within s) \ bounded_linear f' \ (\e>0. \d>0. \x'\s. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" unfolding has_derivative_within Lim_within dist_norm by (simp add: diff_diff_eq) lemma has_derivative_at': "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \x'. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" using has_derivative_within' [of f f' x UNIV] by simp lemma has_derivative_componentwise_within: "(f has_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_derivative (\x. f' x \ i)) (at a within S))" apply (simp add: has_derivative_within) apply (subst tendsto_componentwise_iff) apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib) apply (simp add: algebra_simps) done lemma has_derivative_at_withinI: "(f has_derivative f') (at x) \ (f has_derivative f') (at x within s)" unfolding has_derivative_within' has_derivative_at' by blast lemma has_derivative_right: fixes f :: "real \ real" and y :: "real" shows "(f has_derivative ((*) y)) (at x within ({x <..} \ I)) \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x <..} \ I))" proof - have "((\t. (f t - (f x + y * (t - x))) / \t - x\) \ 0) (at x within ({x<..} \ I)) \ ((\t. (f t - f x) / (t - x) - y) \ 0) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) also have "\ \ ((\t. (f t - f x) / (t - x)) \ y) (at x within ({x<..} \ I))" by (simp add: Lim_null[symmetric]) also have "\ \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (simp_all add: field_simps) finally show ?thesis by (simp add: bounded_linear_mult_right has_derivative_within) qed subsubsection \Caratheodory characterization\ lemma DERIV_caratheodory_within: "(f has_field_derivative l) (at x within S) \ (\g. (\z. f z - f x = g z * (z - x)) \ continuous (at x within S) g \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (intro exI conjI) let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z-x)" by simp show "continuous (at x within S) ?g" using \?lhs\ by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) show "?g x = l" by simp qed next assume ?rhs then obtain g where "(\z. f z - f x = g z * (z-x))" and "continuous (at x within S) g" and "g x = l" by blast thus ?lhs by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) qed subsection \Differentiability\ definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infix "differentiable'_on" 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" lemma differentiableI: "(f has_derivative f') net \ f differentiable net" unfolding differentiable_def by auto lemma differentiable_onD: "\f differentiable_on S; x \ S\ \ f differentiable (at x within S)" using differentiable_on_def by blast lemma differentiable_at_withinI: "f differentiable (at x) \ f differentiable (at x within s)" unfolding differentiable_def using has_derivative_at_withinI by blast lemma differentiable_at_imp_differentiable_on: "(\x. x \ s \ f differentiable at x) \ f differentiable_on s" by (metis differentiable_at_withinI differentiable_on_def) corollary\<^marker>\tag unimportant\ differentiable_iff_scaleR: fixes f :: "real \ 'a::real_normed_vector" shows "f differentiable F \ (\d. (f has_derivative (\x. x *\<^sub>R d)) F)" by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) lemma differentiable_on_eq_differentiable_at: "open s \ f differentiable_on s \ (\x\s. f differentiable at x)" unfolding differentiable_on_def by (metis at_within_interior interior_open) lemma differentiable_transform_within: assumes "f differentiable (at x within s)" and "0 < d" and "x \ s" and "\x'. \x'\s; dist x' x < d\ \ f x' = g x'" shows "g differentiable (at x within s)" using assms has_derivative_transform_within unfolding differentiable_def by blast lemma differentiable_on_ident [simp, derivative_intros]: "(\x. x) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_on_id [simp, derivative_intros]: "id differentiable_on S" by (simp add: id_def) lemma differentiable_on_const [simp, derivative_intros]: "(\z. c) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_mult [simp, derivative_intros]: fixes f :: "'M::real_normed_vector \ 'a::real_normed_algebra" shows "\f differentiable_on S; g differentiable_on S\ \ (\z. f z * g z) differentiable_on S" unfolding differentiable_on_def differentiable_def using differentiable_def differentiable_mult by blast lemma differentiable_on_compose: "\g differentiable_on S; f differentiable_on (g ` S)\ \ (\x. f (g x)) differentiable_on S" by (simp add: differentiable_in_compose differentiable_on_def) lemma bounded_linear_imp_differentiable_on: "bounded_linear f \ f differentiable_on S" by (simp add: differentiable_on_def bounded_linear_imp_differentiable) lemma linear_imp_differentiable_on: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable_on S" by (simp add: differentiable_on_def linear_imp_differentiable) lemma differentiable_on_minus [simp, derivative_intros]: "f differentiable_on S \ (\z. -(f z)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_add [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_diff [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_inverse [simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" shows "f differentiable_on S \ (\x. x \ S \ f x \ 0) \ (\x. inverse (f x)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_scaleR [derivative_intros, simp]: "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" unfolding differentiable_on_def by (blast intro: differentiable_scaleR) lemma has_derivative_sqnorm_at [derivative_intros, simp]: "((\x. (norm x)\<^sup>2) has_derivative (\x. 2 *\<^sub>R (a \ x))) (at a)" using bounded_bilinear.FDERIV [of "(\)" id id a _ id id] by (auto simp: inner_commute dot_square_norm bounded_bilinear_inner) lemma differentiable_sqnorm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "(\x. (norm x)\<^sup>2) differentiable (at a)" by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) lemma differentiable_on_sqnorm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "(\x. (norm x)\<^sup>2) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_norm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "a \ 0 \ norm differentiable (at a)" using differentiableI has_derivative_norm by blast lemma differentiable_on_norm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "0 \ S \ norm differentiable_on S" by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) subsection \Frechet derivative and Jacobian matrix\ definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" proposition frechet_derivative_works: "f differentiable net \ (f has_derivative (frechet_derivative f net)) net" unfolding frechet_derivative_def differentiable_def unfolding some_eq_ex[of "\ f' . (f has_derivative f') net"] .. lemma linear_frechet_derivative: "f differentiable net \ linear (frechet_derivative f net)" unfolding frechet_derivative_works has_derivative_def by (auto intro: bounded_linear.linear) lemma frechet_derivative_const [simp]: "frechet_derivative (\x. c) (at a) = (\x. 0)" using differentiable_const frechet_derivative_works has_derivative_const has_derivative_unique by blast lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id" using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast lemma frechet_derivative_ident [simp]: "frechet_derivative (\x. x) (at a) = (\x. x)" by (metis eq_id_iff frechet_derivative_id) subsection \Differentiability implies continuity\ proposition differentiable_imp_continuous_within: "f differentiable (at x within s) \ continuous (at x within s) f" by (auto simp: differentiable_def intro: has_derivative_continuous) lemma differentiable_imp_continuous_on: "f differentiable_on s \ continuous_on s f" unfolding differentiable_on_def continuous_on_eq_continuous_within using differentiable_imp_continuous_within by blast lemma differentiable_on_subset: "f differentiable_on t \ s \ t \ f differentiable_on s" unfolding differentiable_on_def using differentiable_within_subset by blast lemma differentiable_on_empty: "f differentiable_on {}" unfolding differentiable_on_def by auto lemma has_derivative_continuous_on: "(\x. x \ s \ (f has_derivative f' x) (at x within s)) \ continuous_on s f" by (auto intro!: differentiable_imp_continuous_on differentiableI simp: differentiable_on_def) text \Results about neighborhoods filter.\ lemma eventually_nhds_metric_le: "eventually P (nhds a) = (\d>0. \x. dist x a \ d \ P x)" unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) lemma le_nhds: "F \ nhds a \ (\S. open S \ a \ S \ eventually (\x. x \ S) F)" unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) lemma le_nhds_metric: "F \ nhds a \ (\e>0. eventually (\x. dist x a < e) F)" unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) lemma le_nhds_metric_le: "F \ nhds a \ (\e>0. eventually (\x. dist x a \ e) F)" unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) text \Several results are easier using a "multiplied-out" variant. (I got this idea from Dieudonne's proof of the chain rule).\ lemma has_derivative_within_alt: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. \d>0. \y\s. norm(y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_within_alt2: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. eventually (\y. norm (f y - f x - f' (y - x)) \ e * norm (y - x)) (at x within s))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_at_alt: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \y. norm(y - x) < d \ norm (f y - f x - f'(y - x)) \ e * norm (y - x))" using has_derivative_within_alt[where s=UNIV] by simp subsection \The chain rule\ proposition diff_chain_within[derivative_intros]: assumes "(f has_derivative f') (at x within s)" and "(g has_derivative g') (at (f x) within (f ` s))" shows "((g \ f) has_derivative (g' \ f'))(at x within s)" using has_derivative_in_compose[OF assms] by (simp add: comp_def) lemma diff_chain_at[derivative_intros]: "(f has_derivative f') (at x) \ (g has_derivative g') (at (f x)) \ ((g \ f) has_derivative (g' \ f')) (at x)" using has_derivative_compose[of f f' x UNIV g g'] by (simp add: comp_def) lemma has_vector_derivative_within_open: "a \ S \ open S \ (f has_vector_derivative f') (at a within S) \ (f has_vector_derivative f') (at a)" by (simp only: at_within_interior interior_open) lemma field_vector_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_field_derivative g') (at (f x) within f ` S)" shows "((g \ f) has_vector_derivative (f' * g')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_derivative g') (at (f x) within f`S)" shows "((g \ f) has_vector_derivative (g' f')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg] linear.scaleR[OF has_derivative_linear[OF Dg]] unfolding has_vector_derivative_def o_def by (auto simp: o_def mult.commute has_vector_derivative_def) subsection\<^marker>\tag unimportant\ \Composition rules stated just for differentiability\ lemma differentiable_chain_at: "f differentiable (at x) \ g differentiable (at (f x)) \ (g \ f) differentiable (at x)" unfolding differentiable_def by (meson diff_chain_at) lemma differentiable_chain_within: "f differentiable (at x within S) \ g differentiable (at(f x) within (f ` S)) \ (g \ f) differentiable (at x within S)" unfolding differentiable_def by (meson diff_chain_within) subsection \Uniqueness of derivative\ text\<^marker>\tag important\ \ The general result is a bit messy because we need approachability of the limit point from any direction. But OK for nontrivial intervals etc. \ proposition frechet_derivative_unique_within: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes 1: "(f has_derivative f') (at x within S)" and 2: "(f has_derivative f'') (at x within S)" and S: "\i e. \i\Basis; e>0\ \ \d. 0 < \d\ \ \d\ < e \ (x + d *\<^sub>R i) \ S" shows "f' = f''" proof - note as = assms(1,2)[unfolded has_derivative_def] then interpret f': bounded_linear f' by auto from as interpret f'': bounded_linear f'' by auto have "x islimpt S" unfolding islimpt_approachable proof (intro allI impI) fix e :: real assume "e > 0" obtain d where "0 < \d\" and "\d\ < e" and "x + d *\<^sub>R (SOME i. i \ Basis) \ S" using assms(3) SOME_Basis \e>0\ by blast then show "\x'\S. x' \ x \ dist x' x < e" by (rule_tac x="x + d *\<^sub>R (SOME i. i \ Basis)" in bexI) (auto simp: dist_norm SOME_Basis nonzero_Basis) qed then have *: "netlimit (at x within S) = x" by (simp add: Lim_ident_at trivial_limit_within) show ?thesis proof (rule linear_eq_stdbasis) show "linear f'" "linear f''" unfolding linear_conv_bounded_linear using as by auto next fix i :: 'a assume i: "i \ Basis" define e where "e = norm (f' i - f'' i)" show "f' i = f'' i" proof (rule ccontr) assume "f' i \ f'' i" then have "e > 0" unfolding e_def by auto obtain d where d: "0 < d" "(\y. y\S \ 0 < dist y x \ dist y x < d \ dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) - (f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)" using tendsto_diff [OF as(1,2)[THEN conjunct2]] unfolding * Lim_within using \e>0\ by blast obtain c where c: "0 < \c\" "\c\ < d \ x + c *\<^sub>R i \ S" using assms(3) i d(1) by blast have *: "norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / \c\) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" unfolding scaleR_right_distrib by auto also have "\ = norm ((1 / \c\) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto also have "\ = e" unfolding e_def using c(1) using norm_minus_cancel[of "f' i - f'' i"] by auto finally show False using c using d(2)[of "x + c *\<^sub>R i"] unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib using i by (auto simp: inverse_eq_divide) qed qed qed proposition frechet_derivative_unique_within_closed_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes ab: "\i. i\Basis \ a\i < b\i" and x: "x \ cbox a b" and "(f has_derivative f' ) (at x within cbox a b)" and "(f has_derivative f'') (at x within cbox a b)" shows "f' = f''" proof (rule frechet_derivative_unique_within) fix e :: real fix i :: 'a assume "e > 0" and i: "i \ Basis" then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ cbox a b" proof (cases "x\i = a\i") case True with ab[of i] \e>0\ x i show ?thesis by (rule_tac x="(min (b\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) next case False moreover have "a \ i < x \ i" using False i mem_box(2) x by force moreover { have "a \ i * 2 + min (x \ i - a \ i) e \ a\i *2 + x\i - a\i" by auto also have "\ = a\i + x\i" by auto also have "\ \ 2 * (x\i)" using \a \ i < x \ i\ by auto finally have "a \ i * 2 + min (x \ i - a \ i) e \ x \ i * 2" by auto } moreover have "min (x \ i - a \ i) e \ 0" by (simp add: \0 < e\ \a \ i < x \ i\ less_eq_real_def) then have "x \ i * 2 \ b \ i * 2 + min (x \ i - a \ i) e" using i mem_box(2) x by force ultimately show ?thesis using ab[of i] \e>0\ x i by (rule_tac x="- (min (x\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) qed qed (use assms in auto) lemma frechet_derivative_unique_within_open_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes x: "x \ box a b" and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)" shows "f' = f''" proof - have "at x within box a b = at x" by (metis x at_within_interior interior_open open_box) with f show "f' = f''" by (simp add: has_derivative_unique) qed lemma frechet_derivative_at: "(f has_derivative f') (at x) \ f' = frechet_derivative f (at x)" using differentiable_def frechet_derivative_works has_derivative_unique by blast lemma frechet_derivative_compose: "frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)" if "g differentiable at x" "f differentiable at (g x)" by (metis diff_chain_at frechet_derivative_at frechet_derivative_works that) lemma frechet_derivative_within_cbox: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "\i. i\Basis \ a\i < b\i" and "x \ cbox a b" and "(f has_derivative f') (at x within cbox a b)" shows "frechet_derivative f (at x within cbox a b) = f'" using assms by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) lemma frechet_derivative_transform_within_open: "frechet_derivative f (at x) = frechet_derivative g (at x)" if "f differentiable at x" "open X" "x \ X" "\x. x \ X \ f x = g x" by (meson frechet_derivative_at frechet_derivative_works has_derivative_transform_within_open that) subsection \Derivatives of local minima and maxima are zero\ lemma has_derivative_local_min: fixes f :: "'a::real_normed_vector \ real" assumes deriv: "(f has_derivative f') (at x)" assumes min: "eventually (\y. f x \ f y) (at x)" shows "f' = (\h. 0)" proof fix h :: 'a interpret f': bounded_linear f' using deriv by (rule has_derivative_bounded_linear) show "f' h = 0" proof (cases "h = 0") case False from min obtain d where d1: "0 < d" and d2: "\y\ball x d. f x \ f y" unfolding eventually_at by (force simp: dist_commute) have "FDERIV (\r. x + r *\<^sub>R h) 0 :> (\r. r *\<^sub>R h)" by (intro derivative_eq_intros) auto then have "FDERIV (\r. f (x + r *\<^sub>R h)) 0 :> (\k. f' (k *\<^sub>R h))" by (rule has_derivative_compose, simp add: deriv) then have "DERIV (\r. f (x + r *\<^sub>R h)) 0 :> f' h" unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs) moreover have "0 < d / norm h" using d1 and \h \ 0\ by simp moreover have "\y. \0 - y\ < d / norm h \ f (x + 0 *\<^sub>R h) \ f (x + y *\<^sub>R h)" using \h \ 0\ by (auto simp add: d2 dist_norm pos_less_divide_eq) ultimately show "f' h = 0" by (rule DERIV_local_min) qed simp qed lemma has_derivative_local_max: fixes f :: "'a::real_normed_vector \ real" assumes "(f has_derivative f') (at x)" assumes "eventually (\y. f y \ f x) (at x)" shows "f' = (\h. 0)" using has_derivative_local_min [of "\x. - f x" "\h. - f' h" "x"] using assms unfolding fun_eq_iff by simp lemma differential_zero_maxmin: fixes f::"'a::real_normed_vector \ real" assumes "x \ S" and "open S" and deriv: "(f has_derivative f') (at x)" and mono: "(\y\S. f y \ f x) \ (\y\S. f x \ f y)" shows "f' = (\v. 0)" using mono proof assume "\y\S. f y \ f x" with \x \ S\ and \open S\ have "eventually (\y. f y \ f x) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_max) next assume "\y\S. f x \ f y" with \x \ S\ and \open S\ have "eventually (\y. f x \ f y) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_min) qed lemma differential_zero_maxmin_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" and ball: "0 < e" "(\y \ ball x e. (f y)\k \ (f x)\k) \ (\y\ball x e. (f x)\k \ (f y)\k)" and diff: "f differentiable (at x)" shows "(\j\Basis. (frechet_derivative f (at x) j \ k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") proof - let ?f' = "frechet_derivative f (at x)" have "x \ ball x e" using \0 < e\ by simp moreover have "open (ball x e)" by simp moreover have "((\x. f x \ k) has_derivative (\h. ?f' h \ k)) (at x)" using bounded_linear_inner_left diff[unfolded frechet_derivative_works] by (rule bounded_linear.has_derivative) ultimately have "(\h. frechet_derivative f (at x) h \ k) = (\v. 0)" using ball(2) by (rule differential_zero_maxmin) then show ?thesis unfolding fun_eq_iff by simp qed subsection \One-dimensional mean value theorem\ lemma mvt_simple: fixes f :: "real \ real" assumes "a < b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a<.. real" assumes "a \ b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a..b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" using assms(2) assms(1) by auto case True then show ?thesis by force next case False then show ?thesis using mvt_simple[OF _ derf] by (metis \a \ b\ atLeastAtMost_iff dual_order.order_iff_strict greaterThanLessThan_iff) qed text \A nice generalization (see Havin's proof of 5.19 from Rudin's book).\ lemma mvt_general: fixes f :: "real \ 'a::real_inner" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\x\{a<.. norm (f' x (b - a))" proof - have "\x\{a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" apply (rule mvt [OF \a < b\, where f = "\x. (f b - f a) \ f x"]) apply (intro continuous_intros contf) using derf apply (auto intro: has_derivative_inner_right) done then obtain x where x: "x \ {a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" .. show ?thesis proof (cases "f a = f b") case False have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" by (simp add: power2_eq_square) also have "\ = (f b - f a) \ (f b - f a)" unfolding power2_norm_eq_inner .. also have "\ = (f b - f a) \ f' x (b - a)" using x(2) by (simp only: inner_diff_right) also have "\ \ norm (f b - f a) * norm (f' x (b - a))" by (rule norm_cauchy_schwarz) finally show ?thesis using False x(1) by (auto simp add: mult_left_cancel) next case True then show ?thesis using \a < b\ by (rule_tac x="(a + b) /2" in bexI) auto qed qed subsection \More general bound theorems\ proposition differentiable_bound_general: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and f_cont: "continuous_on {a..b} f" and phi_cont: "continuous_on {a..b} \" and f': "\x. a < x \ x < b \ (f has_vector_derivative f' x) (at x)" and phi': "\x. a < x \ x < b \ (\ has_vector_derivative \' x) (at x)" and bnd: "\x. a < x \ x < b \ norm (f' x) \ \' x" shows "norm (f b - f a) \ \ b - \ a" proof - { fix x assume x: "a < x" "x < b" have "0 \ norm (f' x)" by simp also have "\ \ \' x" using x by (auto intro!: bnd) finally have "0 \ \' x" . } note phi'_nonneg = this note f_tendsto = assms(2)[simplified continuous_on_def, rule_format] note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] { fix e::real assume "e > 0" define e2 where "e2 = e / 2" with \e > 0\ have "e2 > 0" by simp let ?le = "\x1. norm (f x1 - f a) \ \ x1 - \ a + e * (x1 - a) + e" define A where "A = {x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" have A_subset: "A \ {a..b}" by (auto simp: A_def) { fix x2 assume a: "a \ x2" "x2 \ b" and le: "\x1\{a..e > 0\ proof cases assume "x2 \ a" with a have "a < x2" by simp have "at x2 within {a <.. bot" using \a < x2\ by (auto simp: trivial_limit_within islimpt_in_closure) moreover have "((\x1. (\ x1 - \ a) + e * (x1 - a) + e) \ (\ x2 - \ a) + e * (x2 - a) + e) (at x2 within {a <..x1. norm (f x1 - f a)) \ norm (f x2 - f a)) (at x2 within {a <..x. x > a) (at x2 within {a <.. A" using assms by (auto simp: A_def) hence [simp]: "A \ {}" by auto have A_ivl: "\x1 x2. x2 \ A \ x1 \ {a ..x2} \ x1 \ A" by (simp add: A_def) have [simp]: "bdd_above A" by (auto simp: A_def) define y where "y = Sup A" have "y \ b" unfolding y_def by (simp add: cSup_le_iff) (simp add: A_def) have leI: "\x x1. a \ x1 \ x \ A \ x1 < x \ ?le x1" by (auto simp: A_def intro!: le_cont) have y_all_le: "\x1\{a.. y" by (metis \a \ A\ \bdd_above A\ cSup_upper y_def) have "y \ A" using y_all_le \a \ y\ \y \ b\ by (auto simp: A_def) hence "A = {a .. y}" using A_subset by (auto simp: subset_iff y_def cSup_upper intro: A_ivl) from le_cont[OF \a \ y\ \y \ b\ y_all_le] have le_y: "?le y" . have "y = b" proof (cases "a = y") case True with \a < b\ have "y < b" by simp with \a = y\ f_cont phi_cont \e2 > 0\ have 1: "\\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2" and 2: "\\<^sub>F x in at y within {y..b}. dist (\ x) (\ y) < e2" by (auto simp: continuous_on_def tendsto_iff) have 3: "eventually (\x. y < x) (at y within {y..b})" by (auto simp: eventually_at_filter) have 4: "eventually (\x::real. x < b) (at y within {y..b})" using _ \y < b\ by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) from 1 2 3 4 have eventually_le: "eventually (\x. ?le x) (at y within {y .. b})" proof eventually_elim case (elim x1) have "norm (f x1 - f a) = norm (f x1 - f y)" by (simp add: \a = y\) also have "norm (f x1 - f y) \ e2" using elim \a = y\ by (auto simp : dist_norm intro!: less_imp_le) also have "\ \ e2 + (\ x1 - \ a + e2 + e * (x1 - a))" using \0 < e\ elim by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) (auto simp: \a = y\ dist_norm intro!: mult_nonneg_nonneg) also have "\ = \ x1 - \ a + e * (x1 - a) + e" by (simp add: e2_def) finally show "?le x1" . qed from this[unfolded eventually_at_topological] \?le y\ obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y..b} \ ?le x" by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min b (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a = y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" by (simp add: d'_def) fix x1 assume "x1 \ {a.. S" "x1 \ {y..b}" by (auto simp: \a = y\ d'_def dist_real_def intro!: d ) thus "?le x1" by (rule S) qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp then show "y = b" using \d > 0\ \y < b\ by (simp add: d'_def) next case False with \a \ y\ have "a < y" by simp show "y = b" proof (rule ccontr) assume "y \ b" hence "y < b" using \y \ b\ by simp let ?F = "at y within {y.. has_vector_derivative \' y) ?F" using \a < y\ \y < b\ by (auto simp add: at_within_open[of _ "{a<..\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \ e2 * \x1 - y\" "\\<^sub>F x1 in ?F. norm (\ x1 - \ y - (x1 - y) *\<^sub>R \' y) \ e2 * \x1 - y\" using \e2 > 0\ by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) moreover have "\\<^sub>F x1 in ?F. y \ x1" "\\<^sub>F x1 in ?F. x1 < b" by (auto simp: eventually_at_filter) ultimately have "\\<^sub>F x1 in ?F. norm (f x1 - f y) \ (\ x1 - \ y) + e * \x1 - y\" (is "\\<^sub>F x1 in ?F. ?le' x1") proof eventually_elim case (elim x1) from norm_triangle_ineq2[THEN order_trans, OF elim(1)] have "norm (f x1 - f y) \ norm (f' y) * \x1 - y\ + e2 * \x1 - y\" by (simp add: ac_simps) also have "norm (f' y) \ \' y" using bnd \a < y\ \y < b\ by simp also have "\' y * \x1 - y\ \ \ x1 - \ y + e2 * \x1 - y\" using elim by (simp add: ac_simps) finally have "norm (f x1 - f y) \ \ x1 - \ y + e2 * \x1 - y\ + e2 * \x1 - y\" by (auto simp: mult_right_mono) thus ?case by (simp add: e2_def) qed moreover have "?le' y" by simp ultimately obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y.. ?le' x" unfolding eventually_at_topological by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min ((y + b)/2) (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a < y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" using \y < b\ by (simp add: d'_def min_def) fix x1 assume x1: "x1 \ {a..y \ A\ local.leI x1 by auto next case False hence x1': "x1 \ S" "x1 \ {y.. norm (f x1 - f y) + norm (f y - f a)" by (rule order_trans[OF _ norm_triangle_ineq]) simp also note S(3)[OF x1'] also note le_y finally show "?le x1" using False by (auto simp: algebra_simps) qed qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp thus False using \d > 0\ \y < b\ by (simp add: d'_def min_def split: if_split_asm) qed qed with le_y have "norm (f b - f a) \ \ b - \ a + e * (b - a + 1)" by (simp add: algebra_simps) } note * = this show ?thesis proof (rule field_le_epsilon) fix e::real assume "e > 0" then show "norm (f b - f a) \ \ b - \ a + e" using *[of "e / (b - a + 1)"] \a < b\ by simp qed qed lemma differentiable_bound: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\x. x\S \ (f has_derivative f' x) (at x within S)" and B: "\x. x \ S \ onorm (f' x) \ B" and x: "x \ S" and y: "y \ S" shows "norm (f x - f y) \ B * norm (x - y)" proof - let ?p = "\u. x + u *\<^sub>R (y - x)" let ?\ = "\h. h * B * norm (x - y)" have *: "x + u *\<^sub>R (y - x) \ S" if "u \ {0..1}" for u proof - have "u *\<^sub>R y = u *\<^sub>R (y - x) + u *\<^sub>R x" by (simp add: scale_right_diff_distrib) then show "x + u *\<^sub>R (y - x) \ S" using that \convex S\ x y by (simp add: convex_alt) (metis pth_b(2) pth_c(1) scaleR_collapse) qed have "\z. z \ (\u. x + u *\<^sub>R (y - x)) ` {0..1} \ (f has_derivative f' z) (at z within (\u. x + u *\<^sub>R (y - x)) ` {0..1})" by (auto intro: * has_derivative_within_subset [OF derf]) then have "continuous_on (?p ` {0..1}) f" unfolding continuous_on_eq_continuous_within by (meson has_derivative_continuous) with * have 1: "continuous_on {0 .. 1} (f \ ?p)" by (intro continuous_intros)+ { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret linear "(f' ?u)" using u by (auto intro!: has_derivative_linear derf *) have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" by (intro derivative_intros has_derivative_within_subset [OF derf]) (use u * in auto) hence "((f \ ?p) has_vector_derivative f' ?u (y - x)) (at u)" by (simp add: at_within_open[OF u open_greaterThanLessThan] scaleR has_vector_derivative_def o_def) } note 2 = this have 3: "continuous_on {0..1} ?\" by (rule continuous_intros)+ have 4: "(?\ has_vector_derivative B * norm (x - y)) (at u)" for u by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret bounded_linear "(f' ?u)" using u by (auto intro!: has_derivative_bounded_linear derf *) have "norm (f' ?u (y - x)) \ onorm (f' ?u) * norm (y - x)" by (rule onorm) (rule bounded_linear) also have "onorm (f' ?u) \ B" using u by (auto intro!: assms(3)[rule_format] *) finally have "norm ((f' ?u) (y - x)) \ B * norm (x - y)" by (simp add: mult_right_mono norm_minus_commute) } note 5 = this have "norm (f x - f y) = norm ((f \ (\u. x + u *\<^sub>R (y - x))) 1 - (f \ (\u. x + u *\<^sub>R (y - x))) 0)" by (auto simp add: norm_minus_commute) also from differentiable_bound_general[OF zero_less_one 1, OF 3 2 4 5] have "norm ((f \ ?p) 1 - (f \ ?p) 0) \ B * norm (x - y)" by simp finally show ?thesis . qed lemma differentiable_bound_segment: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes "\t. t \ {0..1} \ x0 + t *\<^sub>R a \ G" assumes f': "\x. x \ G \ (f has_derivative f' x) (at x within G)" assumes B: "\x. x \ {0..1} \ onorm (f' (x0 + x *\<^sub>R a)) \ B" shows "norm (f (x0 + a) - f x0) \ norm a * B" proof - let ?G = "(\x. x0 + x *\<^sub>R a) ` {0..1}" have "?G = (+) x0 ` (\x. x *\<^sub>R a) ` {0..1}" by auto also have "convex \" by (intro convex_translation convex_scaled convex_real_interval) finally have "convex ?G" . moreover have "?G \ G" "x0 \ ?G" "x0 + a \ ?G" using assms by (auto intro: image_eqI[where x=1]) ultimately show ?thesis using has_derivative_subset[OF f' \?G \ G\] B differentiable_bound[of "(\x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0] by (force simp: ac_simps) qed lemma differentiable_bound_linearization: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes S: "\t. t \ {0..1} \ a + t *\<^sub>R (b - a) \ S" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes B: "\x. x \ S \ onorm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - f' x0 (b - a)) \ norm (b - a) * B" proof - define g where [abs_def]: "g x = f x - f' x0 x" for x have g: "\x. x \ S \ (g has_derivative (\i. f' x i - f' x0 i)) (at x within S)" unfolding g_def using assms by (auto intro!: derivative_eq_intros bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f']) from B have "\x\{0..1}. onorm (\i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \ B" using assms by (auto simp: fun_diff_def) with differentiable_bound_segment[OF S g] \x0 \ S\ show ?thesis by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']]) qed lemma vector_differentiable_bound_linearization: fixes f::"real \ 'b::real_normed_vector" assumes f': "\x. x \ S \ (f has_vector_derivative f' x) (at x within S)" assumes "closed_segment a b \ S" assumes B: "\x. x \ S \ norm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - (b - a) *\<^sub>R f' x0) \ norm (b - a) * B" using assms by (intro differentiable_bound_linearization[of a b S f "\x h. h *\<^sub>R f' x" x0 B]) (force simp: closed_segment_real_eq has_vector_derivative_def scaleR_diff_right[symmetric] mult.commute[of B] intro!: onorm_le mult_left_mono)+ text \In particular.\ lemma has_derivative_zero_constant: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" shows "\c. \x\s. f x = c" proof - { fix x y assume "x \ s" "y \ s" then have "norm (f x - f y) \ 0 * norm (x - y)" using assms by (intro differentiable_bound[of s]) (auto simp: onorm_zero) then have "f x = f y" by simp } then show ?thesis by metis qed lemma has_field_derivative_zero_constant: assumes "convex s" "\x. x \ s \ (f has_field_derivative 0) (at x within s)" shows "\c. \x\s. f (x) = (c :: 'a :: real_normed_field)" proof (rule has_derivative_zero_constant) have A: "(*) 0 = (\_. 0 :: 'a)" by (intro ext) simp fix x assume "x \ s" thus "(f has_derivative (\h. 0)) (at x within s)" using assms(2)[of x] by (simp add: has_field_derivative_def A) qed fact lemma has_vector_derivative_zero_constant: assumes "convex s" assumes "\x. x \ s \ (f has_vector_derivative 0) (at x within s)" obtains c where "\x. x \ s \ f x = c" using has_derivative_zero_constant[of s f] assms by (auto simp: has_vector_derivative_def) lemma has_derivative_zero_unique: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" and "x \ s" "y \ s" shows "f x = f y" using has_derivative_zero_constant[OF assms(1,2)] assms(3-) by force lemma has_derivative_zero_unique_connected: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open s" "connected s" assumes f: "\x. x \ s \ (f has_derivative (\x. 0)) (at x)" assumes "x \ s" "y \ s" shows "f x = f y" proof (rule connected_local_const[where f=f, OF \connected s\ \x\s\ \y\s\]) show "\a\s. eventually (\b. f a = f b) (at a within s)" proof fix a assume "a \ s" with \open s\ obtain e where "0 < e" "ball a e \ s" by (rule openE) then have "\c. \x\ball a e. f x = c" by (intro has_derivative_zero_constant) (auto simp: at_within_open[OF _ open_ball] f convex_ball) with \0 have "\x\ball a e. f a = f x" by auto then show "eventually (\b. f a = f b) (at a within s)" using \0 unfolding eventually_at_topological by (intro exI[of _ "ball a e"]) auto qed qed subsection \Differentiability of inverse function (most basic form)\ lemma has_derivative_inverse_basic: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes derf: "(f has_derivative f') (at (g y))" and ling': "bounded_linear g'" and "g' \ f' = id" and contg: "continuous (at y) g" and "open T" and "y \ T" and fg: "\z. z \ T \ f (g z) = z" shows "(g has_derivative g') (at y)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain C where C: "0 < C" "\x. norm (g' x) \ norm x * C" using bounded_linear.pos_bounded[OF assms(2)] by blast have lem1: "\e>0. \d>0. \z. norm (z - y) < d \ norm (g z - g y - g'(z - y)) \ e * norm (g z - g y)" proof (intro allI impI) fix e :: real assume "e > 0" with C(1) have *: "e / C > 0" by auto obtain d0 where "0 < d0" and d0: "\u. norm (u - g y) < d0 \ norm (f u - f (g y) - f' (u - g y)) \ e / C * norm (u - g y)" using derf * unfolding has_derivative_at_alt by blast obtain d1 where "0 < d1" and d1: "\x. \0 < dist x y; dist x y < d1\ \ dist (g x) (g y) < d0" using contg \0 < d0\ unfolding continuous_at Lim_at by blast obtain d2 where "0 < d2" and d2: "\u. dist u y < d2 \ u \ T" using \open T\ \y \ T\ unfolding open_dist by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF \0 < d1\ \0 < d2\] by blast show "\d>0. \z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < d" then have "z \ T" using d2 d unfolding dist_norm by auto have "norm (g z - g y - g' (z - y)) \ norm (g' (f (g z) - y - f' (g z - g y)))" unfolding g'.diff f'.diff unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] fg[OF \z\T\] by (simp add: norm_minus_commute) also have "\ \ norm (f (g z) - y - f' (g z - g y)) * C" by (rule C(2)) also have "\ \ (e / C) * norm (g z - g y) * C" proof - have "norm (g z - g y) < d0" by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \0 < d0\ d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff) then show ?thesis by (metis C(1) \y \ T\ d0 fg real_mult_le_cancel_iff1) qed also have "\ \ e * norm (g z - g y)" using C by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" by simp qed (use d in auto) qed have *: "(0::real) < 1 / 2" by auto obtain d where "0 < d" and d: "\z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ 1/2 * norm (g z - g y)" using lem1 * by blast define B where "B = C * 2" have "B > 0" unfolding B_def using C by auto have lem2: "norm (g z - g y) \ B * norm (z - y)" if z: "norm(z - y) < d" for z proof - have "norm (g z - g y) \ norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" by (rule norm_triangle_sub) also have "\ \ norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" by (rule add_left_mono) (use d z in auto) also have "\ \ norm (z - y) * C + 1 / 2 * norm (g z - g y)" by (rule add_right_mono) (use C in auto) finally show "norm (g z - g y) \ B * norm (z - y)" unfolding B_def by (auto simp add: field_simps) qed show ?thesis unfolding has_derivative_at_alt proof (intro conjI assms allI impI) fix e :: real assume "e > 0" then have *: "e / B > 0" by (metis \B > 0\ divide_pos_pos) obtain d' where "0 < d'" and d': "\z. norm (z - y) < d' \ norm (g z - g y - g' (z - y)) \ e / B * norm (g z - g y)" using lem1 * by blast obtain k where k: "0 < k" "k < d" "k < d'" using field_lbound_gt_zero[OF \0 < d\ \0 < d'\] by blast show "\d>0. \ya. norm (ya - y) < d \ norm (g ya - g y - g' (ya - y)) \ e * norm (ya - y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < k" then have "norm (g z - g y - g' (z - y)) \ e / B * norm(g z - g y)" using d' k by auto also have "\ \ e * norm (z - y)" unfolding times_divide_eq_left pos_divide_le_eq[OF \B>0\] using lem2[of z] k as \e > 0\ by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (z - y)" by simp qed (use k in auto) qed qed text \Simply rewrite that based on the domain point x.\ lemma has_derivative_inverse_basic_x: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" and "continuous (at (f x)) g" and "g (f x) = x" and "open T" and "f x \ T" and "\y. y \ T \ f (g y) = y" shows "(g has_derivative g') (at (f x))" by (rule has_derivative_inverse_basic) (use assms in auto) text \This is the version in Dieudonne', assuming continuity of f and g.\ lemma has_derivative_inverse_dieudonne: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open S" and "open (f ` S)" and "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g (f x) = x" and "x \ S" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) using assms(3-6) unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] apply auto done text \Here's the simplest way of not assuming much about g.\ proposition has_derivative_inverse: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "compact S" and "x \ S" and fx: "f x \ interior (f ` S)" and "continuous_on S f" and gf: "\y. y \ S \ g (f y) = y" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" proof - have *: "\y. y \ interior (f ` S) \ f (g y) = y" by (metis gf image_iff interior_subset subsetCE) show ?thesis apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"]) apply (rule continuous_on_interior[OF _ fx]) apply (rule continuous_on_inv) apply (simp_all add: assms *) done qed text \Invertible derivative continuous at a point implies local injectivity. It's only for this we need continuity of the derivative, except of course if we want the fact that the inverse derivative is also continuous. So if we know for some other reason that the inverse function exists, it's OK.\ proposition has_derivative_locally_injective: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "a \ S" and "open S" and bling: "bounded_linear g'" and "g' \ f' a = id" and derf: "\x. x \ S \ (f has_derivative f' x) (at x)" and "\e. e > 0 \ \d>0. \x. dist a x < d \ onorm (\v. f' x v - f' a v) < e" obtains r where "r > 0" "ball a r \ S" "inj_on f (ball a r)" proof - interpret bounded_linear g' using assms by auto note f'g' = assms(4)[unfolded id_def o_def,THEN cong] have "g' (f' a (\Basis)) = (\Basis)" "(\Basis) \ (0::'n)" using f'g' by auto then have *: "0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)] by fastforce define k where "k = 1 / onorm g' / 2" have *: "k > 0" unfolding k_def using * by auto obtain d1 where d1: "0 < d1" "\x. dist a x < d1 \ onorm (\v. f' x v - f' a v) < k" using assms(6) * by blast from \open S\ obtain d2 where "d2 > 0" "ball a d2 \ S" using \a\S\ .. obtain d2 where d2: "0 < d2" "ball a d2 \ S" using \0 < d2\ \ball a d2 \ S\ by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF d1(1) d2(1)] by blast show ?thesis proof show "0 < d" by (fact d) show "ball a d \ S" using \d < d2\ \ball a d2 \ S\ by auto show "inj_on f (ball a d)" unfolding inj_on_def proof (intro strip) fix x y assume as: "x \ ball a d" "y \ ball a d" "f x = f y" define ph where [abs_def]: "ph w = w - g' (f w - f x)" for w have ph':"ph = g' \ (\w. f' a w - (f w - f x))" unfolding ph_def o_def by (simp add: diff f'g') have "norm (ph x - ph y) \ (1 / 2) * norm (x - y)" proof (rule differentiable_bound[OF convex_ball _ _ as(1-2)]) fix u assume u: "u \ ball a d" then have "u \ S" using d d2 by auto have *: "(\v. v - g' (f' u v)) = g' \ (\w. f' a w - f' u w)" unfolding o_def and diff using f'g' by auto have blin: "bounded_linear (f' a)" using \a \ S\ derf by blast show "(ph has_derivative (\v. v - g' (f' u v))) (at u within ball a d)" unfolding ph' * comp_def by (rule \u \ S\ derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin] bounded_linear.has_derivative [OF bling] |simp)+ have **: "bounded_linear (\x. f' u x - f' a x)" "bounded_linear (\x. f' a x - f' u x)" using \u \ S\ blin bounded_linear_sub derf by auto then have "onorm (\v. v - g' (f' u v)) \ onorm g' * onorm (\w. f' a w - f' u w)" by (simp add: "*" bounded_linear_axioms onorm_compose) also have "\ \ onorm g' * k" apply (rule mult_left_mono) using d1(2)[of u] using onorm_neg[where f="\x. f' u x - f' a x"] d u onorm_pos_le[OF bling] apply (auto simp: algebra_simps) done also have "\ \ 1 / 2" unfolding k_def by auto finally show "onorm (\v. v - g' (f' u v)) \ 1 / 2" . qed moreover have "norm (ph y - ph x) = norm (y - x)" by (simp add: as(3) ph_def) ultimately show "x = y" unfolding norm_minus_commute by auto qed qed qed subsection \Uniformly convergent sequence of derivatives\ lemma has_derivative_sequence_lipschitz_lemma: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\n x h. \n\N; x \ S\ \ norm (f' n x h - g' x h) \ e * norm h" and "0 \ e" shows "\m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof clarify fix m n x y assume as: "N \ m" "N \ n" "x \ S" "y \ S" show "norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof (rule differentiable_bound[where f'="\x h. f' m x h - f' n x h", OF \convex S\ _ _ as(3-4)]) fix x assume "x \ S" show "((\a. f m a - f n a) has_derivative (\h. f' m x h - f' n x h)) (at x within S)" by (rule derivative_intros derf \x\S\)+ show "onorm (\h. f' m x h - f' n x h) \ 2 * e" proof (rule onorm_bound) fix h have "norm (f' m x h - f' n x h) \ norm (f' m x h - g' x h) + norm (f' n x h - g' x h)" using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] by (auto simp add: algebra_simps norm_minus_commute) also have "\ \ e * norm h + e * norm h" using nle[OF \N \ m\ \x \ S\, of h] nle[OF \N \ n\ \x \ S\, of h] by (auto simp add: field_simps) finally show "norm (f' m x h - f' n x h) \ 2 * e * norm h" by auto qed (simp add: \0 \ e\) qed qed lemma has_derivative_sequence_Lipschitz: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "e > 0" shows "\N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" proof - have *: "2 * (e/2) = e" using \e > 0\ by auto obtain N where "\n\N. \x\S. \h. norm (f' n x h - g' x h) \ (e/2) * norm h" using nle \e > 0\ unfolding eventually_sequentially by (metis less_divide_eq_numeral1(1) mult_zero_left) then show "\N. \m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" apply (rule_tac x=N in exI) apply (rule has_derivative_sequence_lipschitz_lemma[where e="e/2", unfolded *]) using assms \e > 0\ apply auto done qed proposition has_derivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "x0 \ S" and lim: "((\n. f n x0) \ l) sequentially" shows "\g. \x\S. (\n. f n x) \ g x \ (g has_derivative g'(x)) (at x within S)" proof - have lem1: "\e. e > 0 \ \N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" using assms(1,2,3) by (rule has_derivative_sequence_Lipschitz) have "\g. \x\S. ((\n. f n x) \ g x) sequentially" proof (intro ballI bchoice) fix x assume "x \ S" show "\y. (\n. f n x) \ y" unfolding convergent_eq_Cauchy proof (cases "x = x0") case True then show "Cauchy (\n. f n x)" using LIMSEQ_imp_Cauchy[OF lim] by auto next case False show "Cauchy (\n. f n x)" unfolding Cauchy_def proof (intro allI impI) fix e :: real assume "e > 0" hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto obtain M where M: "\m\M. \n\M. dist (f m x0) (f n x0) < e / 2" using LIMSEQ_imp_Cauchy[OF lim] * unfolding Cauchy_def by blast obtain N where N: "\m\N. \n\N. \u\S. \y\S. norm (f m u - f n u - (f m y - f n y)) \ e / 2 / norm (x - x0) * norm (u - y)" using lem1 *(2) by blast show "\M. \m\M. \n\M. dist (f m x) (f n x) < e" proof (intro exI allI impI) fix m n assume as: "max M N \m" "max M N\n" have "dist (f m x) (f n x) \ norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" unfolding dist_norm by (rule norm_triangle_sub) also have "\ \ norm (f m x0 - f n x0) + e / 2" using N \x\S\ \x0\S\ as False by fastforce also have "\ < e / 2 + e / 2" by (rule add_strict_right_mono) (use as M in \auto simp: dist_norm\) finally show "dist (f m x) (f n x) < e" by auto qed qed qed qed then obtain g where g: "\x\S. (\n. f n x) \ g x" .. have lem2: "\N. \n\N. \x\S. \y\S. norm ((f n x - f n y) - (g x - g y)) \ e * norm (x - y)" if "e > 0" for e proof - obtain N where N: "\m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" using lem1 \e > 0\ by blast show "\N. \n\N. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" proof (intro exI ballI allI impI) fix n x y assume as: "N \ n" "x \ S" "y \ S" have "((\m. norm (f n x - f n y - (f m x - f m y))) \ norm (f n x - f n y - (g x - g y))) sequentially" by (intro tendsto_intros g[rule_format] as) moreover have "eventually (\m. norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)) sequentially" unfolding eventually_sequentially proof (intro exI allI impI) fix m assume "N \ m" then show "norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)" using N as by (auto simp add: algebra_simps) qed ultimately show "norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" by (simp add: tendsto_upperbound) qed qed have "\x\S. ((\n. f n x) \ g x) sequentially \ (g has_derivative g' x) (at x within S)" unfolding has_derivative_within_alt2 proof (intro ballI conjI allI impI) fix x assume "x \ S" then show "(\n. f n x) \ g x" by (simp add: g) have tog': "(\n. f' n x u) \ g' x u" for u unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm proof (intro allI impI) fix e :: real assume "e > 0" show "eventually (\n. norm (f' n x u - g' x u) \ e) sequentially" proof (cases "u = 0") case True have "eventually (\n. norm (f' n x u - g' x u) \ e * norm u) sequentially" using nle \0 < e\ \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u = 0\ \0 < e\ by (auto elim: eventually_mono) next case False with \0 < e\ have "0 < e / norm u" by simp then have "eventually (\n. norm (f' n x u - g' x u) \ e / norm u * norm u) sequentially" using nle \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u \ 0\ by simp qed qed show "bounded_linear (g' x)" proof fix x' y z :: 'a fix c :: real note lin = assms(2)[rule_format,OF \x\S\,THEN has_derivative_bounded_linear] show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] apply (intro tendsto_intros tog') done show "g' x (y + z) = g' x y + g' x z" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_add] apply (rule tendsto_add) apply (rule tog')+ done obtain N where N: "\h. norm (f' N x h - g' x h) \ 1 * norm h" using nle \x \ S\ unfolding eventually_sequentially by (fast intro: zero_less_one) have "bounded_linear (f' N x)" using derf \x \ S\ by fast from bounded_linear.bounded [OF this] obtain K where K: "\h. norm (f' N x h) \ norm h * K" .. { fix h have "norm (g' x h) = norm (f' N x h - (f' N x h - g' x h))" by simp also have "\ \ norm (f' N x h) + norm (f' N x h - g' x h)" by (rule norm_triangle_ineq4) also have "\ \ norm h * K + 1 * norm h" using N K by (fast intro: add_mono) finally have "norm (g' x h) \ norm h * (K + 1)" by (simp add: ring_distribs) } then show "\K. \h. norm (g' x h) \ norm h * K" by fast qed show "eventually (\y. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)) (at x within S)" if "e > 0" for e proof - have *: "e / 3 > 0" using that by auto obtain N1 where N1: "\n\N1. \x\S. \h. norm (f' n x h - g' x h) \ e / 3 * norm h" using nle * unfolding eventually_sequentially by blast obtain N2 where N2[rule_format]: "\n\N2. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e / 3 * norm (x - y)" using lem2 * by blast let ?N = "max N1 N2" have "eventually (\y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)) (at x within S)" using derf[unfolded has_derivative_within_alt2] and \x \ S\ and * by fast moreover have "eventually (\y. y \ S) (at x within S)" unfolding eventually_at by (fast intro: zero_less_one) ultimately show "\\<^sub>F y in at x within S. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" proof (rule eventually_elim2) fix y assume "y \ S" assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)" moreover have "norm (g y - g x - (f ?N y - f ?N x)) \ e / 3 * norm (y - x)" using N2[OF _ \y \ S\ \x \ S\] by (simp add: norm_minus_commute) ultimately have "norm (g y - g x - f' ?N x (y - x)) \ 2 * e / 3 * norm (y - x)" using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] by (auto simp add: algebra_simps) moreover have " norm (f' ?N x (y - x) - g' x (y - x)) \ e / 3 * norm (y - x)" using N1 \x \ S\ by auto ultimately show "norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by (auto simp add: algebra_simps) qed qed qed then show ?thesis by fast qed text \Can choose to line up antiderivatives if we want.\ lemma has_antiderivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and der: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and no: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof (cases "S = {}") case False then obtain a where "a \ S" by auto have *: "\P Q. \g. \x\S. P g x \ Q g x \ \g. \x\S. Q g x" by auto show ?thesis apply (rule *) apply (rule has_derivative_sequence [OF \convex S\ _ no, of "\n x. f n x + (f 0 a - f n a)"]) apply (metis assms(2) has_derivative_add_const) using \a \ S\ apply auto done qed auto lemma has_antiderivative_limit: fixes g' :: "'a::real_normed_vector \ 'a \ 'b::banach" assumes "convex S" and "\e. e>0 \ \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm (f' x h - g' x h) \ e * norm h)" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof - have *: "\n. \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm(f' x h - g' x h) \ inverse (real (Suc n)) * norm h)" by (simp add: assms(2)) obtain f where *: "\x. \f'. \xa\S. (f x has_derivative f' xa) (at xa within S) \ (\h. norm (f' xa h - g' xa h) \ inverse (real (Suc x)) * norm h)" using * by metis obtain f' where f': "\x. \z\S. (f x has_derivative f' x z) (at z within S) \ (\h. norm (f' x z h - g' z h) \ inverse (real (Suc x)) * norm h)" using * by metis show ?thesis proof (rule has_antiderivative_sequence[OF \convex S\, of f f']) fix e :: real assume "e > 0" obtain N where N: "inverse (real (Suc N)) < e" using reals_Archimedean[OF \e>0\] .. show "\\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" unfolding eventually_sequentially proof (intro exI allI ballI impI) fix n x h assume n: "N \ n" and x: "x \ S" have *: "inverse (real (Suc n)) \ e" apply (rule order_trans[OF _ N[THEN less_imp_le]]) using n apply (auto simp add: field_simps) done show "norm (f' n x h - g' x h) \ e * norm h" by (meson "*" mult_right_mono norm_ge_zero order.trans x f') qed qed (use f' in auto) qed subsection \Differentiation of a series\ proposition has_derivative_series: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (sum (\i. f' i x h) {.. e * norm h" and "x \ S" and "(\n. f n x) sums l" shows "\g. \x\S. (\n. f n x) sums (g x) \ (g has_derivative g' x) (at x within S)" unfolding sums_def apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) apply (metis assms(2) has_derivative_sum) using assms(4-5) unfolding sums_def apply auto done lemma has_field_derivative_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniform_limit S (\n x. \i S" "summable (\n. f n x0)" shows "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" unfolding has_field_derivative_def proof (rule has_derivative_series) show "\\<^sub>F n in sequentially. \x\S. \h. norm ((\i e * norm h" if "e > 0" for e unfolding eventually_sequentially proof - from that assms(3) obtain N where N: "\n x. n \ N \ x \ S \ norm ((\i N" "x \ S" have "norm ((\iii e" by simp hence "norm ((\i e * norm h" by (intro mult_right_mono) simp_all finally have "norm ((\i e * norm h" . } thus "\N. \n\N. \x\S. \h. norm ((\i e * norm h" by blast qed qed (use assms in \auto simp: has_field_derivative_def\) lemma has_field_derivative_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" "x \ interior S" shows "summable (\n. f n x)" "((\x. \n. f n x) has_field_derivative (\n. f' n x)) (at x)" proof - from \x \ interior S\ have "x \ S" using interior_subset by blast define g' where [abs_def]: "g' x = (\i. f' i x)" for x from assms(3) have "uniform_limit S (\n x. \ix. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(1)[OF \x \ S\] show "summable (\n. f n x)" by (simp add: sums_iff) from g(2)[OF \x \ S\] \x \ interior S\ have "(g has_field_derivative g' x) (at x)" by (simp add: at_within_interior[of x S]) also have "(g has_field_derivative g' x) (at x) \ ((\x. \n. f n x) has_field_derivative g' x) (at x)" using eventually_nhds_in_nhd[OF \x \ interior S\] interior_subset[of S] g(1) by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) finally show "((\x. \n. f n x) has_field_derivative g' x) (at x)" . qed lemma differentiable_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "summable (\n. f n x)" and "(\x. \n. f n x) differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g[OF x] show "summable (\n. f n x)" by (auto simp: summable_def) from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def differentiable_def has_field_derivative_def) qed lemma differentiable_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" shows "(\x. \n. f n x) differentiable (at x0)" using differentiable_series[OF assms, of x0] \x0 \ S\ by blast+ subsection \Derivative as a vector\ text \Considering derivative \<^typ>\real \ 'b::real_normed_vector\ as a vector.\ definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" lemma vector_derivative_unique_within: assumes not_bot: "at x within S \ bot" and f': "(f has_vector_derivative f') (at x within S)" and f'': "(f has_vector_derivative f'') (at x within S)" shows "f' = f''" proof - have "(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" proof (rule frechet_derivative_unique_within, simp_all) show "\d. d \ 0 \ \d\ < e \ x + d \ S" if "0 < e" for e proof - from that obtain x' where "x' \ S" "x' \ x" "\x' - x\ < e" using islimpt_approachable_real[of x S] not_bot by (auto simp add: trivial_limit_within) then show ?thesis using eq_iff_diff_eq_0 by fastforce qed qed (use f' f'' in \auto simp: has_vector_derivative_def\) then show ?thesis unfolding fun_eq_iff by (metis scaleR_one) qed lemma vector_derivative_unique_at: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f'') (at x) \ f' = f''" by (rule vector_derivative_unique_within) auto lemma differentiableI_vector: "(f has_vector_derivative y) F \ f differentiable F" by (auto simp: differentiable_def has_vector_derivative_def) proposition vector_derivative_works: "f differentiable net \ (f has_vector_derivative (vector_derivative f net)) net" (is "?l = ?r") proof assume ?l obtain f' where f': "(f has_derivative f') net" using \?l\ unfolding differentiable_def .. then interpret bounded_linear f' by auto show ?r unfolding vector_derivative_def has_vector_derivative_def by (rule someI[of _ "f' 1"]) (simp add: scaleR[symmetric] f') qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def) lemma vector_derivative_within: assumes not_bot: "at x within S \ bot" and y: "(f has_vector_derivative y) (at x within S)" shows "vector_derivative f (at x within S) = y" using y by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y]) (auto simp: differentiable_def has_vector_derivative_def) lemma frechet_derivative_eq_vector_derivative: assumes "f differentiable (at x)" shows "(frechet_derivative f (at x)) = (\r. r *\<^sub>R vector_derivative f (at x))" using assms by (auto simp: differentiable_iff_scaleR vector_derivative_def has_vector_derivative_def intro: someI frechet_derivative_at [symmetric]) lemma has_real_derivative: fixes f :: "real \ real" assumes "(f has_derivative f') F" obtains c where "(f has_real_derivative c) F" proof - obtain c where "f' = (\x. x * c)" by (metis assms has_derivative_bounded_linear real_bounded_linear) then show ?thesis by (metis assms that has_field_derivative_def mult_commute_abs) qed lemma has_real_derivative_iff: fixes f :: "real \ real" shows "(\c. (f has_real_derivative c) F) = (\D. (f has_derivative D) F)" by (metis has_field_derivative_def has_real_derivative) lemma has_vector_derivative_cong_ev: assumes *: "eventually (\x. x \ S \ f x = g x) (nhds x)" "f x = g x" shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def has_derivative_def using * apply (cases "at x within S \ bot") apply (intro refl conj_cong filterlim_cong) apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono) done lemma islimpt_closure_open: fixes s :: "'a::perfect_space set" assumes "open s" and t: "t = closure s" "x \ t" shows "x islimpt t" proof cases assume "x \ s" { fix T assume "x \ T" "open T" then have "open (s \ T)" using \open s\ by auto then have "s \ T \ {x}" using not_open_singleton[of x] by auto with \x \ T\ \x \ s\ have "\y\t. y \ T \ y \ x" using closure_subset[of s] by (auto simp: t) } then show ?thesis by (auto intro!: islimptI) next assume "x \ s" with t show ?thesis unfolding t closure_def by (auto intro: islimpt_subset) qed lemma vector_derivative_unique_within_closed_interval: assumes ab: "a < b" "x \ cbox a b" assumes D: "(f has_vector_derivative f') (at x within cbox a b)" "(f has_vector_derivative f'') (at x within cbox a b)" shows "f' = f''" using ab by (intro vector_derivative_unique_within[OF _ D]) (auto simp: trivial_limit_within intro!: islimpt_closure_open[where s="{a <..< b}"]) lemma vector_derivative_at: "(f has_vector_derivative f') (at x) \ vector_derivative f (at x) = f'" by (intro vector_derivative_within at_neq_bot) lemma has_vector_derivative_id_at [simp]: "vector_derivative (\x. x) (at a) = 1" by (simp add: vector_derivative_at) lemma vector_derivative_minus_at [simp]: "f differentiable at a \ vector_derivative (\x. - f x) (at a) = - vector_derivative f (at a)" by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric]) lemma vector_derivative_add_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric]) lemma vector_derivative_diff_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) lemma vector_derivative_mult_at [simp]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) lemma vector_derivative_scaleR_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" apply (rule vector_derivative_at) apply (rule has_vector_derivative_scaleR) apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) done lemma vector_derivative_within_cbox: assumes ab: "a < b" "x \ cbox a b" assumes f: "(f has_vector_derivative f') (at x within cbox a b)" shows "vector_derivative f (at x within cbox a b) = f'" by (intro vector_derivative_unique_within_closed_interval[OF ab _ f] vector_derivative_works[THEN iffD1] differentiableI_vector) fact lemma vector_derivative_within_closed_interval: fixes f::"real \ 'a::euclidean_space" assumes "a < b" and "x \ {a..b}" assumes "(f has_vector_derivative f') (at x within {a..b})" shows "vector_derivative f (at x within {a..b}) = f'" using assms vector_derivative_within_cbox by fastforce lemma has_vector_derivative_within_subset: "(f has_vector_derivative f') (at x within S) \ T \ S \ (f has_vector_derivative f') (at x within T)" by (auto simp: has_vector_derivative_def intro: has_derivative_within_subset) lemma has_vector_derivative_at_within: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def by (rule has_derivative_at_withinI) lemma has_vector_derivative_weaken: fixes x D and f g S T assumes f: "(f has_vector_derivative D) (at x within T)" and "x \ S" "S \ T" and "\x. x \ S \ f x = g x" shows "(g has_vector_derivative D) (at x within S)" proof - have "(f has_vector_derivative D) (at x within S) \ (g has_vector_derivative D) (at x within S)" unfolding has_vector_derivative_def has_derivative_iff_norm using assms by (intro conj_cong Lim_cong_within refl) auto then show ?thesis using has_vector_derivative_within_subset[OF f \S \ T\] by simp qed lemma has_vector_derivative_transform_within: assumes "(f has_vector_derivative f') (at x within S)" and "0 < d" and "x \ S" and "\x'. \x'\S; dist x' x < d\ \ f x' = g x'" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within) lemma has_vector_derivative_transform_within_open: assumes "(f has_vector_derivative f') (at x)" and "open S" and "x \ S" and "\y. y\S \ f y = g y" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within_open) lemma has_vector_derivative_transform: assumes "x \ S" "\x. x \ S \ g x = f x" assumes f': "(f has_vector_derivative f') (at x within S)" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform) lemma vector_diff_chain_at: assumes "(f has_vector_derivative f') (at x)" and "(g has_vector_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x)" using assms has_vector_derivative_at_within has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_diff_chain_within: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at (f x) within f ` s)" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" using assms has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_derivative_const_at [simp]: "vector_derivative (\x. c) (at a) = 0" by (simp add: vector_derivative_at) lemma vector_derivative_at_within_ivl: "(f has_vector_derivative f') (at x) \ a \ x \ x \ b \ a vector_derivative f (at x within {a..b}) = f'" using has_vector_derivative_at_within vector_derivative_within_cbox by fastforce lemma vector_derivative_chain_at: assumes "f differentiable at x" "(g differentiable at (f x))" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) assumes Df: "(f has_vector_derivative f') (at x)" and Dg: "(g has_field_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' * g')) (at x)" using diff_chain_at[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_chain_within: assumes "at x within S \ bot" "f differentiable (at x within S)" "(g has_derivative g') (at (f x) within f ` S)" shows "vector_derivative (g \ f) (at x within S) = g' (vector_derivative f (at x within S)) " apply (rule vector_derivative_within [OF \at x within S \ bot\]) apply (rule vector_derivative_diff_chain_within) using assms(2-3) vector_derivative_works by auto subsection \Field differentiability\ definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" (infixr "(field'_differentiable)" 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" lemma field_differentiable_imp_differentiable: "f field_differentiable F \ f differentiable F" unfolding field_differentiable_def differentiable_def using has_field_derivative_imp_has_derivative by auto lemma field_differentiable_imp_continuous_at: "f field_differentiable (at x within S) \ continuous (at x within S) f" by (metis DERIV_continuous field_differentiable_def) lemma field_differentiable_within_subset: "\f field_differentiable (at x within S); T \ S\ \ f field_differentiable (at x within T)" by (metis DERIV_subset field_differentiable_def) lemma field_differentiable_at_within: "\f field_differentiable (at x)\ \ f field_differentiable (at x within S)" unfolding field_differentiable_def by (metis DERIV_subset top_greatest) lemma field_differentiable_linear [simp,derivative_intros]: "((*) c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def mult_commute_abs by (force intro: has_derivative_mult_right) lemma field_differentiable_const [simp,derivative_intros]: "(\z. c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_const has_field_derivative_imp_has_derivative by blast lemma field_differentiable_ident [simp,derivative_intros]: "(\z. z) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_ident has_field_derivative_def by blast lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F" unfolding id_def by (rule field_differentiable_ident) lemma field_differentiable_minus [derivative_intros]: "f field_differentiable F \ (\z. - (f z)) field_differentiable F" unfolding field_differentiable_def by (metis field_differentiable_minus) lemma field_differentiable_add [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z + g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_add) lemma field_differentiable_add_const [simp,derivative_intros]: "(+) c field_differentiable F" by (simp add: field_differentiable_add) lemma field_differentiable_sum [derivative_intros]: "(\i. i \ I \ (f i) field_differentiable F) \ (\z. \i\I. f i z) field_differentiable F" by (induct I rule: infinite_finite_induct) (auto intro: field_differentiable_add field_differentiable_const) lemma field_differentiable_diff [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z - g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_diff) lemma field_differentiable_inverse [derivative_intros]: assumes "f field_differentiable (at a within S)" "f a \ 0" shows "(\z. inverse (f z)) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_inverse_fun) lemma field_differentiable_mult [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" shows "(\z. f z * g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_mult [of f _ a S g]) lemma field_differentiable_divide [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" "g a \ 0" shows "(\z. f z / g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_divide [of f _ a S g]) lemma field_differentiable_power [derivative_intros]: assumes "f field_differentiable (at a within S)" shows "(\z. f z ^ n) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_power) lemma field_differentiable_transform_within: "0 < d \ x \ S \ (\x'. x' \ S \ dist x' x < d \ f x' = g x') \ f field_differentiable (at x within S) \ g field_differentiable (at x within S)" unfolding field_differentiable_def has_field_derivative_def by (blast intro: has_derivative_transform_within) lemma field_differentiable_compose_within: assumes "f field_differentiable (at a within S)" "g field_differentiable (at (f a) within f`S)" shows "(g o f) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_image_chain) lemma field_differentiable_compose: "f field_differentiable at z \ g field_differentiable at (f z) \ (g o f) field_differentiable at z" by (metis field_differentiable_at_within field_differentiable_compose_within) lemma field_differentiable_within_open: "\a \ S; open S\ \ f field_differentiable at a within S \ f field_differentiable at a" unfolding field_differentiable_def by (metis at_within_open) lemma exp_scaleR_has_vector_derivative_right: "((\t. exp (t *\<^sub>R A)) has_vector_derivative exp (t *\<^sub>R A) * A) (at t within T)" unfolding has_vector_derivative_def proof (rule has_derivativeI) let ?F = "at t within (T \ {t - 1 <..< t + 1})" have *: "at t within T = ?F" by (rule at_within_nhd[where S="{t - 1 <..< t + 1}"]) auto let ?e = "\i x. (inverse (1 + real i) * inverse (fact i) * (x - t) ^ i) *\<^sub>R (A * A ^ i)" have "\\<^sub>F n in sequentially. \x\T \ {t - 1<.. norm (A ^ (n + 1) /\<^sub>R fact (n + 1))" apply (auto simp: algebra_split_simps intro!: eventuallyI) apply (rule mult_left_mono) apply (auto simp add: field_simps power_abs intro!: divide_right_mono power_le_one) done then have "uniform_limit (T \ {t - 1<..n x. \ix. \i. ?e i x) sequentially" by (rule Weierstrass_m_test_ev) (intro summable_ignore_initial_segment summable_norm_exp) moreover have "\\<^sub>F x in sequentially. x > 0" by (metis eventually_gt_at_top) then have "\\<^sub>F n in sequentially. ((\x. \i A) ?F" by eventually_elim (auto intro!: tendsto_eq_intros simp: power_0_left if_distrib if_distribR cong: if_cong) ultimately have [tendsto_intros]: "((\x. \i. ?e i x) \ A) ?F" by (auto intro!: swap_uniform_limit[where f="\n x. \i < n. ?e i x" and F = sequentially]) have [tendsto_intros]: "((\x. if x = t then 0 else 1) \ 1) ?F" by (rule tendsto_eventually) (simp add: eventually_at_filter) have "((\y. ((y - t) / abs (y - t)) *\<^sub>R ((\n. ?e n y) - A)) \ 0) (at t within T)" unfolding * by (rule tendsto_norm_zero_cancel) (auto intro!: tendsto_eq_intros) moreover have "\\<^sub>F x in at t within T. x \ t" by (simp add: eventually_at_filter) then have "\\<^sub>F x in at t within T. ((x - t) / \x - t\) *\<^sub>R ((\n. ?e n x) - A) = (exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" proof eventually_elim case (elim x) have "(exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = ((\n. (x - t) *\<^sub>R ?e n x) - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" unfolding exp_first_term by (simp add: ac_simps) also have "summable (\n. ?e n x)" proof - from elim have "?e n x = (((x - t) *\<^sub>R A) ^ (n + 1)) /\<^sub>R fact (n + 1) /\<^sub>R (x - t)" for n by simp then show ?thesis by (auto simp only: intro!: summable_scaleR_right summable_ignore_initial_segment summable_exp_generic) qed then have "(\n. (x - t) *\<^sub>R ?e n x) = (x - t) *\<^sub>R (\n. ?e n x)" by (rule suminf_scaleR_right[symmetric]) also have "(\ - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = (x - t) *\<^sub>R ((\n. ?e n x) - A) /\<^sub>R norm (x - t)" by (simp add: algebra_simps) finally show ?case by simp (simp add: field_simps) qed ultimately have "((\y. (exp ((y - t) *\<^sub>R A) - 1 - (y - t) *\<^sub>R A) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) from tendsto_mult_right_zero[OF this, where c="exp (t *\<^sub>R A)"] show "((\y. (exp (y *\<^sub>R A) - exp (t *\<^sub>R A) - (y - t) *\<^sub>R (exp (t *\<^sub>R A) * A)) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) (auto simp: algebra_simps field_split_simps exp_add_commuting[symmetric]) qed (rule bounded_linear_scaleR_left) lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)" using exp_times_arg_commute[symmetric, of "t *\<^sub>R A"] by (auto simp: algebra_simps) lemma exp_scaleR_has_vector_derivative_left: "((\t. exp (t *\<^sub>R A)) has_vector_derivative A * exp (t *\<^sub>R A)) (at t)" using exp_scaleR_has_vector_derivative_right[of A t] by (simp add: exp_times_scaleR_commute) subsection \Field derivative\ definition\<^marker>\tag important\ deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where "deriv f x \ SOME D. DERIV f x :> D" lemma DERIV_imp_deriv: "DERIV f x :> f' \ deriv f x = f'" unfolding deriv_def by (metis some_equality DERIV_unique) lemma DERIV_deriv_iff_has_field_derivative: "DERIV f x :> deriv f x \ (\f'. (f has_field_derivative f') (at x))" by (auto simp: has_field_derivative_def DERIV_imp_deriv) lemma DERIV_deriv_iff_real_differentiable: fixes x :: real shows "DERIV f x :> deriv f x \ f differentiable at x" unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) lemma deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "deriv f x = deriv g y" proof - have "(\D. (f has_field_derivative D) (at x)) = (\D. (g has_field_derivative D) (at y))" by (intro ext DERIV_cong_ev refl assms) thus ?thesis by (simp add: deriv_def assms) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma real_derivative_chain: fixes x :: real shows "f differentiable at x \ g differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv) lemma field_derivative_eq_vector_derivative: "(deriv f x) = vector_derivative f (at x)" by (simp add: mult.commute deriv_def vector_derivative_def has_vector_derivative_def has_field_derivative_def) proposition field_differentiable_derivI: "f field_differentiable (at x) \ (f has_field_derivative deriv f x) (at x)" by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) lemma vector_derivative_chain_at_general: assumes "f differentiable at x" "g field_differentiable at (f x)" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) * deriv g (f x)" apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) using assms vector_derivative_works by (auto simp: field_differentiable_derivI) subsection \Relation between convexity and derivative\ (* TODO: Generalise to real vector spaces? *) proposition convex_on_imp_above_tangent: assumes convex: "convex_on A f" and connected: "connected A" assumes c: "c \ interior A" and x : "x \ A" assumes deriv: "(f has_field_derivative f') (at c within A)" shows "f x - f c \ f' * (x - c)" proof (cases x c rule: linorder_cases) assume xc: "x > c" let ?A' = "interior A \ {c<..}" from c have "c \ interior A \ closure {c<..}" by auto also have "\ \ closure (interior A \ {c<..})" by (intro open_Int_closure_subset) auto finally have "at c within ?A' \ bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_right_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_right c)" proof eventually_elim fix y assume y: "y \ {c<.. (f x - f c) / (x - c) * (y - c) + f c" using interior_subset[of A] by (intro convex_onD_Icc' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) / (x - c) * (y - c)" by simp thus "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_upperbound) thus ?thesis using xc by (simp add: field_simps) next assume xc: "x < c" let ?A' = "interior A \ {.. interior A \ closure {.. \ closure (interior A \ {.. bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_left_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_left c)" proof eventually_elim fix y assume y: "y \ {x<.. (f x - f c) / (c - x) * (c - y) + f c" using interior_subset[of A] by (intro convex_onD_Icc'' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) * ((c - y) / (c - x))" by simp also have "(c - y) / (c - x) = (y - c) / (x - c)" using y xc by (simp add: field_simps) finally show "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_lowerbound) thus ?thesis using xc by (simp add: field_simps) qed simp_all subsection \Partial derivatives\ lemma eventually_at_Pair_within_TimesI1: fixes x::"'a::metric_space" assumes "\\<^sub>F x' in at x within X. P x'" assumes "P x" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "x \ S" "\x'. x' \ X \ x' \ S \ P x'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" unfolding eventually_at_topological by (auto intro!: exI[where x="S \ UNIV"] S open_Times) qed lemma eventually_at_Pair_within_TimesI2: fixes x::"'a::metric_space" assumes "\\<^sub>F y' in at y within Y. P y'" "P y" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "y \ S" "\y'. y' \ Y \ y' \ S \ P y'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" unfolding eventually_at_topological by (auto intro!: exI[where x="UNIV \ S"] S open_Times) qed proposition has_derivative_partialsI: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" assumes fx: "((\x. f x y) has_derivative fx) (at x within X)" assumes fy: "\x y. x \ X \ y \ Y \ ((\y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" assumes fy_cont[unfolded continuous_within]: "continuous (at (x, y) within X \ Y) (\(x, y). fy x y)" assumes "y \ Y" "convex Y" shows "((\(x, y). f x y) has_derivative (\(tx, ty). fx tx + fy x y ty)) (at (x, y) within X \ Y)" proof (safe intro!: has_derivativeI tendstoI, goal_cases) case (2 e') interpret fx: bounded_linear "fx" using fx by (rule has_derivative_bounded_linear) define e where "e = e' / 9" have "e > 0" using \e' > 0\ by (simp add: e_def) from fy_cont[THEN tendstoD, OF \e > 0\] have "\\<^sub>F (x', y') in at (x, y) within X \ Y. dist (fy x' y') (fy x y) < e" by (auto simp: split_beta') from this[unfolded eventually_at] obtain d' where "d' > 0" "\x' y'. x' \ X \ y' \ Y \ (x', y') \ (x, y) \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" by auto then have d': "x' \ X \ y' \ Y \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" for x' y' using \0 < e\ by (cases "(x', y') = (x, y)") auto define d where "d = d' / sqrt 2" have "d > 0" using \0 < d'\ by (simp add: d_def) have d: "x' \ X \ y' \ Y \ dist x' x < d \ dist y' y < d \ dist (fy x' y') (fy x y) < e" for x' y' by (auto simp: dist_prod_def d_def intro!: d' real_sqrt_sum_squares_less) let ?S = "ball y d \ Y" have "convex ?S" by (auto intro!: convex_Int \convex Y\) { fix x'::'a and y'::'b assume x': "x' \ X" and y': "y' \ Y" assume dx': "dist x' x < d" and dy': "dist y' y < d" have "norm (fy x' y' - fy x' y) \ dist (fy x' y') (fy x y) + dist (fy x' y) (fy x y)" by norm also have "dist (fy x' y') (fy x y) < e" by (rule d; fact) also have "dist (fy x' y) (fy x y) < e" by (auto intro!: d simp: dist_prod_def x' \d > 0\ \y \ Y\ dx') finally have "norm (fy x' y' - fy x' y) < e + e" by arith then have "onorm (blinfun_apply (fy x' y') - blinfun_apply (fy x' y)) < e + e" by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) } note onorm = this have ev_mem: "\\<^sub>F (x', y') in at (x, y) within X \ Y. (x', y') \ X \ Y" using \y \ Y\ by (auto simp: eventually_at intro!: zero_less_one) moreover have ev_dist: "\\<^sub>F xy in at (x, y) within X \ Y. dist xy (x, y) < d" if "d > 0" for d using eventually_at_ball[OF that] by (rule eventually_elim2) (auto simp: dist_commute mem_ball intro!: eventually_True) note ev_dist[OF \0 < d\] ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" proof (eventually_elim, safe) fix x' y' assume "x' \ X" and y': "y' \ Y" assume dist: "dist (x', y') (x, y) < d" then have dx: "dist x' x < d" and dy: "dist y' y < d" unfolding dist_prod_def fst_conv snd_conv atomize_conj by (metis le_less_trans real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) { fix t::real assume "t \ {0 .. 1}" then have "y + t *\<^sub>R (y' - y) \ closed_segment y y'" by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) also have "\ \ ball y d \ Y" using \y \ Y\ \0 < d\ dy y' by (intro \convex ?S\[unfolded convex_contains_segment, rule_format, of y y']) (auto simp: dist_commute) finally have "y + t *\<^sub>R (y' - y) \ ?S" . } note seg = this have "\x. x \ ball y d \ Y \ onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \ e + e" by (safe intro!: onorm less_imp_le \x' \ X\ dx) (auto simp: dist_commute \0 < d\ \y \ Y\) with seg has_derivative_within_subset[OF assms(2)[OF \x' \ X\]] show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" by (rule differentiable_bound_linearization[where S="?S"]) (auto intro!: \0 < d\ \y \ Y\) qed moreover let ?le = "\x'. norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" from fx[unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. ?le x'" by eventually_elim (simp, simp add: dist_norm field_split_simps split: if_split_asm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. ?le x'" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps) moreover have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((x', y') - (x, y)) \ 0" unfolding norm_eq_zero right_minus_eq by (auto simp: eventually_at intro!: zero_less_one) moreover from fy_cont[THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" unfolding eventually_at using \y \ Y\ by (auto simp: dist_prod_def dist_norm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (fy x' y - fy x y) < e" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps \0 < e\) ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" apply eventually_elim proof safe fix x' y' have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \ norm (f x' y' - f x' y - fy x' y (y' - y)) + norm (fy x y (y' - y) - fy x' y (y' - y)) + norm (f x' y - f x y - fx (x' - x))" by norm also assume nz: "norm ((x', y') - (x, y)) \ 0" and nfy: "norm (fy x' y - fy x y) < e" assume "norm (f x' y' - f x' y - blinfun_apply (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" also assume "norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" also have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \ norm ((fy x y) - (fy x' y)) * norm (y' - y)" by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) also have "\ \ (e + e) * norm (y' - y)" using \e > 0\ nfy by (auto simp: norm_minus_commute intro!: mult_right_mono) also have "norm (x' - x) * e \ norm (x' - x) * (e + e)" using \0 < e\ by simp also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \ (norm (y' - y) + norm (x' - x)) * (4 * e)" using \e > 0\ by (simp add: algebra_simps) also have "\ \ 2 * norm ((x', y') - (x, y)) * (4 * e)" using \0 < e\ real_sqrt_sum_squares_ge1[of "norm (x' - x)" "norm (y' - y)"] real_sqrt_sum_squares_ge2[of "norm (y' - y)" "norm (x' - x)"] by (auto intro!: mult_right_mono simp: norm_prod_def simp del: real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) also have "\ \ norm ((x', y') - (x, y)) * (8 * e)" by simp also have "\ < norm ((x', y') - (x, y)) * e'" using \0 < e'\ nz by (auto simp: e_def) finally show "norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" by (simp add: dist_norm) (auto simp add: field_split_simps) qed then show ?case by eventually_elim (auto simp: dist_norm field_simps) next from has_derivative_bounded_linear[OF fx] obtain fxb where "fx = blinfun_apply fxb" by (metis bounded_linear_Blinfun_apply) then show "bounded_linear (\(tx, ty). fx tx + blinfun_apply (fy x y) ty)" by (auto intro!: bounded_linear_intros simp: split_beta') qed subsection\<^marker>\tag unimportant\ \Differentiable case distinction\ lemma has_derivative_within_If_eq: "((\x. if P x then f x else g x) has_derivative f') (at x within S) = (bounded_linear f' \ ((\y.(if P y then (f y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x) else (g y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x))) \ 0) (at x within S))" (is "_ = (_ \ (?if \ 0) _)") proof - have "(\y. (1 / norm (y - x)) *\<^sub>R ((if P y then f y else g y) - ((if P x then f x else g x) + f' (y - x)))) = ?if" by (auto simp: inverse_eq_divide) thus ?thesis by (auto simp: has_derivative_within) qed lemma has_derivative_If_within_closures: assumes f': "x \ S \ (closure S \ closure T) \ (f has_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" assumes x_in: "x \ S \ T" shows "((\x. if x \ S then f x else g x) has_derivative (if x \ S then f' x else g' x)) (at x within (S \ T))" proof - from f' x_in interpret f': bounded_linear "if x \ S then f' x else (\x. 0)" by (auto simp add: has_derivative_within) from g' interpret g': bounded_linear "if x \ T then g' x else (\x. 0)" by (auto simp add: has_derivative_within) have bl: "bounded_linear (if x \ S then f' x else g' x)" using f'.scaleR f'.bounded f'.add g'.scaleR g'.bounded g'.add x_in by (unfold_locales; force) show ?thesis using f' g' closure_subset[of T] closure_subset[of S] unfolding has_derivative_within_If_eq by (intro conjI bl tendsto_If_within_closures x_in) (auto simp: has_derivative_within inverse_eq_divide connect connect' subsetD) qed lemma has_vector_derivative_If_within_closures: assumes x_in: "x \ S \ T" assumes "u = S \ T" assumes f': "x \ S \ (closure S \ closure T) \ (f has_vector_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_vector_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" shows "((\x. if x \ S then f x else g x) has_vector_derivative (if x \ S then f' x else g' x)) (at x within u)" unfolding has_vector_derivative_def assms using x_in apply (intro has_derivative_If_within_closures[where ?f' = "\x a. a *\<^sub>R f' x" and ?g' = "\x a. a *\<^sub>R g' x", THEN has_derivative_eq_rhs]) subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption) subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption) by (auto simp: assms) subsection\<^marker>\tag important\\The Inverse Function Theorem\ lemma linear_injective_contraction: assumes "linear f" "c < 1" and le: "\x. norm (f x - x) \ c * norm x" shows "inj f" unfolding linear_injective_0[OF \linear f\] proof safe fix x assume "f x = 0" with le [of x] have "norm x \ c * norm x" by simp then show "x = 0" using \c < 1\ by (simp add: mult_le_cancel_right1) qed text\From an online proof by J. Michael Boardman, Department of Mathematics, Johns Hopkins University\ lemma inverse_function_theorem_scaled: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative blinfun_apply (f' x)) (at x)" and contf: "continuous_on U f'" and "0 \ U" and [simp]: "f 0 = 0" and id: "f' 0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "0 \ U'" "open V" "0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - obtain d1 where "cball 0 d1 \ U" "d1 > 0" using \open U\ \0 \ U\ open_contains_cball by blast obtain d2 where d2: "\x. \x \ U; dist x 0 \ d2\ \ dist (f' x) (f' 0) < 1/2" "0 < d2" using continuous_onE [OF contf, of 0 "1/2"] by (metis \0 \ U\ half_gt_zero_iff zero_less_one) obtain \ where le: "\x. norm x \ \ \ dist (f' x) id_blinfun \ 1/2" and "0 < \" and subU: "cball 0 \ \ U" proof show "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) show "cball 0 (min d1 d2) \ U" using \cball 0 d1 \ U\ by auto show "dist (f' x) id_blinfun \ 1/2" if "norm x \ min d1 d2" for x using \cball 0 d1 \ U\ d2 that id by fastforce qed let ?D = "cball 0 \" define V:: "'a set" where "V \ ball 0 (\/2)" have 4: "norm (f (x + h) - f x - h) \ 1/2 * norm h" if "x \ ?D" "x+h \ ?D" for x h proof - let ?w = "\x. f x - x" have B: "\x. x \ ?D \ onorm (blinfun_apply (f' x - id_blinfun)) \ 1/2" by (metis dist_norm le mem_cball_0 norm_blinfun.rep_eq) have "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x)" by (rule derivative_eq_intros derf subsetD [OF subU] | force simp: blinfun.diff_left)+ then have Dw: "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x within ?D)" using has_derivative_at_withinI by blast have "norm (?w (x+h) - ?w x) \ (1/2) * norm h" using differentiable_bound [OF convex_cball Dw B] that by fastforce then show ?thesis by (auto simp: algebra_simps) qed have for_g: "\!x. norm x < \ \ f x = y" if y: "norm y < \/2" for y proof - let ?u = "\x. x + (y - f x)" have *: "norm (?u x) < \" if "x \ ?D" for x proof - have fxx: "norm (f x - x) \ \/2" using 4 [of 0 x] \0 < \\ \f 0 = 0\ that by auto have "norm (?u x) \ norm y + norm (f x - x)" by (metis add.commute add_diff_eq norm_minus_commute norm_triangle_ineq) also have "\ < \/2 + \/2" using fxx y by auto finally show ?thesis by simp qed have "\!x \ ?D. ?u x = x" proof (rule banach_fix) show "cball 0 \ \ {}" using \0 < \\ by auto show "(\x. x + (y - f x)) ` cball 0 \ \ cball 0 \" using * by force have "dist (x + (y - f x)) (xh + (y - f xh)) * 2 \ dist x xh" if "norm x \ \" and "norm xh \ \" for x xh using that 4 [of x "xh-x"] by (auto simp: dist_norm norm_minus_commute algebra_simps) then show "\x\cball 0 \. \ya\cball 0 \. dist (x + (y - f x)) (ya + (y - f ya)) \ (1/2) * dist x ya" by auto qed (auto simp: complete_eq_closed) then show ?thesis by (metis "*" add_cancel_right_right eq_iff_diff_eq_0 le_less mem_cball_0) qed define g where "g \ \y. THE x. norm x < \ \ f x = y" have g: "norm (g y) < \ \ f (g y) = y" if "norm y < \/2" for y unfolding g_def using that theI' [OF for_g] by meson then have fg[simp]: "f (g y) = y" if "y \ V" for y using that by (auto simp: V_def) have 5: "norm (g y' - g y) \ 2 * norm (y' - y)" if "y \ V" "y' \ V" for y y' proof - have no: "norm (g y) \ \" "norm (g y') \ \" and [simp]: "f (g y) = y" using that g unfolding V_def by force+ have "norm (g y' - g y) \ norm (g y' - g y - (y' - y)) + norm (y' - y)" by (simp add: add.commute norm_triangle_sub) also have "\ \ (1/2) * norm (g y' - g y) + norm (y' - y)" using 4 [of "g y" "g y' - g y"] that no by (simp add: g norm_minus_commute V_def) finally show ?thesis by auto qed have contg: "continuous_on V g" proof fix y::'a and e::real assume "0 < e" and y: "y \ V" show "\d>0. \x'\V. dist x' y < d \ dist (g x') (g y) \ e" proof (intro exI conjI ballI impI) show "0 < e/2" by (simp add: \0 < e\) qed (use 5 y in \force simp: dist_norm\) qed show thesis proof define U' where "U' \ (f -` V) \ ball 0 \" have contf: "continuous_on U f" using derf has_derivative_at_withinI by (fast intro: has_derivative_continuous_on) then have "continuous_on (ball 0 \) f" by (meson ball_subset_cball continuous_on_subset subU) then show "open U'" by (simp add: U'_def V_def Int_commute continuous_open_preimage) show "0 \ U'" "U' \ U" "open V" "0 \ V" using \0 < \\ subU by (auto simp: U'_def V_def) show hom: "homeomorphism U' V f g" proof show "continuous_on U' f" using \U' \ U\ contf continuous_on_subset by blast show "continuous_on V g" using contg by blast show "f ` U' \ V" using U'_def by blast show "g ` V \ U'" by (simp add: U'_def V_def g image_subset_iff) show "g (f x) = x" if "x \ U'" for x by (metis that fg Int_iff U'_def V_def for_g g mem_ball_0 vimage_eq) show "f (g y) = y" if "y \ V" for y using that by (simp add: g V_def) qed show bij: "bij (blinfun_apply (f'(g y)))" if "y \ V" for y proof - have inj: "inj (blinfun_apply (f' (g y)))" proof (rule linear_injective_contraction) show "linear (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by blast next fix x have "norm (blinfun_apply (f' (g y)) x - x) = norm (blinfun_apply (f' (g y) - id_blinfun) x)" by (simp add: blinfun.diff_left) also have "\ \ norm (f' (g y) - id_blinfun) * norm x" by (rule norm_blinfun) also have "\ \ (1/2) * norm x" proof (rule mult_right_mono) show "norm (f' (g y) - id_blinfun) \ 1/2" using that g [of y] le by (auto simp: V_def dist_norm) qed auto finally show "norm (blinfun_apply (f' (g y)) x - x) \ (1/2) * norm x" . qed auto moreover have "surj (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by (blast intro!: linear_inj_imp_surj [OF _ inj]) ultimately show ?thesis using bijI by blast qed define g' where "g' \ \y. inv (blinfun_apply (f'(g y)))" show "(g has_derivative g' y) (at y)" if "y \ V" for y proof - have gy: "g y \ U" using g subU that unfolding V_def by fastforce obtain e where e: "\h. f (g y + h) = y + blinfun_apply (f' (g y)) h + e h" and e0: "(\h. norm (e h) / norm h) \0\ 0" using iffD1 [OF has_derivative_iff_Ex derf [OF gy]] \y \ V\ by auto have [simp]: "e 0 = 0" using e [of 0] that by simp let ?INV = "inv (blinfun_apply (f' (g y)))" have inj: "inj (blinfun_apply (f' (g y)))" using bij bij_betw_def that by blast have "(g has_derivative g' y) (at y within V)" unfolding has_derivative_at_within_iff_Ex [OF \y \ V\ \open V\] proof show blinv: "bounded_linear (g' y)" unfolding g'_def using derf gy inj inj_linear_imp_inv_bounded_linear by blast define eg where "eg \ \k. - ?INV (e (g (y+k) - g y))" have "g (y+k) = g y + g' y k + eg k" if "y + k \ V" for k proof - have "?INV k = ?INV (blinfun_apply (f' (g y)) (g (y+k) - g y) + e (g (y+k) - g y))" using e [of "g(y+k) - g y"] that by simp then have "g (y+k) = g y + ?INV k - ?INV (e (g (y+k) - g y))" using inj blinv by (simp add: linear_simps g'_def) then show ?thesis by (auto simp: eg_def g'_def) qed moreover have "(\k. norm (eg k) / norm k) \0\ 0" proof (rule Lim_null_comparison) let ?g = "\k. 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" show "\\<^sub>F k in at 0. norm (norm (eg k) / norm k) \ ?g k" unfolding eventually_at_topological proof (intro exI conjI ballI impI) show "open ((+)(-y) ` V)" using \open V\ open_translation by blast show "0 \ (+)(-y) ` V" by (simp add: that) show "norm (norm (eg k) / norm k) \ 2 * onorm (inv (blinfun_apply (f' (g y)))) * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" if "k \ (+)(-y) ` V" "k \ 0" for k proof - have "y+k \ V" using that by auto have "norm (norm (eg k) / norm k) \ onorm ?INV * norm (e (g (y+k) - g y)) / norm k" using blinv g'_def onorm by (force simp: eg_def divide_simps) also have "\ = (norm (g (y+k) - g y) / norm k) * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" by (simp add: divide_simps) also have "\ \ 2 * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" apply (rule mult_right_mono) using 5 [of y "y+k"] \y \ V\ \y + k \ V\ onorm_pos_le [OF blinv] apply (auto simp: divide_simps zero_le_mult_iff zero_le_divide_iff g'_def) done finally show "norm (norm (eg k) / norm k) \ 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" by simp qed qed have 1: "(\h. norm (e h) / norm h) \0\ (norm (e 0) / norm 0)" using e0 by auto have 2: "(\k. g (y+k) - g y) \0\ 0" using contg \open V\ \y \ V\ LIM_offset_zero_iff LIM_zero_iff at_within_open continuous_on_def by fastforce from tendsto_compose [OF 1 2, simplified] have "(\k. norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)) \0\ 0" . from tendsto_mult_left [OF this] show "?g \0\ 0" by auto qed ultimately show "\e. (\k. y + k \ V \ g (y+k) = g y + g' y k + e k) \ (\k. norm (e k) / norm k) \0\ 0" by blast qed then show ?thesis by (metis \open V\ at_within_open that) qed show "g' y = inv (blinfun_apply (f' (g y)))" if "y \ V" for y by (simp add: g'_def) qed qed text\We need all this to justify the scaling and translations.\ theorem inverse_function_theorem: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative (blinfun_apply (f' x))) (at x)" and contf: "continuous_on U f'" and "x0 \ U" and invf: "invf o\<^sub>L f' x0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "x0 \ U'" "open V" "f x0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - have apply1 [simp]: "\i. blinfun_apply invf (blinfun_apply (f' x0) i) = i" by (metis blinfun_apply_blinfun_compose blinfun_apply_id_blinfun invf) have apply2 [simp]: "\i. blinfun_apply (f' x0) (blinfun_apply invf i) = i" by (metis apply1 bij_inv_eq_iff blinfun_bij1 invf) have [simp]: "(range (blinfun_apply invf)) = UNIV" using apply1 surjI by blast let ?f = "invf \ (\x. (f \ (+)x0)x - f x0)" let ?f' = "\x. invf o\<^sub>L (f' (x + x0))" obtain U' V g g' where "open U'" and U': "U' \ (+)(-x0) ` U" "0 \ U'" and "open V" "0 \ V" and hom: "homeomorphism U' V ?f g" and derg: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv (?f'(g y))" and bij: "\y. y \ V \ bij (?f'(g y))" proof (rule inverse_function_theorem_scaled [of "(+)(-x0) ` U" ?f "?f'"]) show ope: "open ((+) (- x0) ` U)" using \open U\ open_translation by blast show "(?f has_derivative blinfun_apply (?f' x)) (at x)" if "x \ (+) (- x0) ` U" for x using that apply clarify apply (rule derf derivative_eq_intros | simp add: blinfun_compose.rep_eq)+ done have YY: "(\x. f' (x + x0)) \u-x0\ f' u" if "f' \u\ f' u" "u \ U" for u using that LIM_offset [where k = x0] by (auto simp: algebra_simps) then have "continuous_on ((+) (- x0) ` U) (\x. f' (x + x0))" using contf \open U\ Lim_at_imp_Lim_at_within by (fastforce simp: continuous_on_def at_within_open_NO_MATCH ope) then show "continuous_on ((+) (- x0) ` U) ?f'" by (intro continuous_intros) simp qed (auto simp: invf \x0 \ U\) show thesis proof let ?U' = "(+)x0 ` U'" let ?V = "((+)(f x0) \ f' x0) ` V" let ?g = "(+)x0 \ g \ invf \ (+)(- f x0)" let ?g' = "\y. inv (blinfun_apply (f' (?g y)))" show oU': "open ?U'" by (simp add: \open U'\ open_translation) show subU: "?U' \ U" using ComplI \U' \ (+) (- x0) ` U\ by auto show "x0 \ ?U'" by (simp add: \0 \ U'\) show "open ?V" using blinfun_bij2 [OF invf] by (metis \open V\ bij_is_surj blinfun.bounded_linear_right bounded_linear_def image_comp open_surjective_linear_image open_translation) show "f x0 \ ?V" using \0 \ V\ image_iff by fastforce show "homeomorphism ?U' ?V f ?g" proof show "continuous_on ?U' f" by (meson subU continuous_on_eq_continuous_at derf has_derivative_continuous oU' subsetD) have "?f ` U' \ V" using hom homeomorphism_image1 by blast then show "f ` ?U' \ ?V" unfolding image_subset_iff by (clarsimp simp: image_def) (metis apply2 add.commute diff_add_cancel) show "?g ` ?V \ ?U'" using hom invf by (auto simp: image_def homeomorphism_def) show "?g (f x) = x" if "x \ ?U'" for x using that hom homeomorphism_apply1 by fastforce have "continuous_on V g" using hom homeomorphism_def by blast then show "continuous_on ?V ?g" by (intro continuous_intros) (auto elim!: continuous_on_subset) have fg: "?f (g x) = x" if "x \ V" for x using hom homeomorphism_apply2 that by blast show "f (?g y) = y" if "y \ ?V" for y using that fg by (simp add: image_iff) (metis apply2 add.commute diff_add_cancel) qed show "(?g has_derivative ?g' y) (at y)" "bij (blinfun_apply (f' (?g y)))" if "y \ ?V" for y proof - have 1: "bij (blinfun_apply invf)" using blinfun_bij1 invf by blast then have 2: "bij (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x by (metis add.commute bij bij_betw_comp_iff2 blinfun_compose.rep_eq that top_greatest) then show "bij (blinfun_apply (f' (?g y)))" using that by auto have "g' x \ blinfun_apply invf = inv (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x using that by (simp add: g' o_inv_distrib blinfun_compose.rep_eq 1 2 add.commute bij_is_inj flip: o_assoc) then show "(?g has_derivative ?g' y) (at y)" using that invf by clarsimp (rule derg derivative_eq_intros | simp flip: id_def)+ qed qed auto qed end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,7684 +1,7079 @@ (* Title: HOL/Analysis/Starlike.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 *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Abstract_Limits begin -section \Line Segments\ - -subsection \Midpoint\ - -definition\<^marker>\tag important\ midpoint :: "'a::real_vector \ 'a \ 'a" - where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" - -lemma midpoint_idem [simp]: "midpoint x x = x" - unfolding midpoint_def by simp - -lemma midpoint_sym: "midpoint a b = midpoint b a" - unfolding midpoint_def by (auto simp add: scaleR_right_distrib) - -lemma midpoint_eq_iff: "midpoint a b = c \ a + b = c + c" -proof - - have "midpoint a b = c \ scaleR 2 (midpoint a b) = scaleR 2 c" - by simp - then show ?thesis - unfolding midpoint_def scaleR_2 [symmetric] by simp -qed - -lemma - fixes a::real - assumes "a \ b" shows ge_midpoint_1: "a \ midpoint a b" - and le_midpoint_1: "midpoint a b \ b" - by (simp_all add: midpoint_def assms) - -lemma dist_midpoint: - fixes a b :: "'a::real_normed_vector" shows - "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) - "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) - "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) - "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) -proof - - have *: "\x y::'a. 2 *\<^sub>R x = - y \ norm x = (norm y) / 2" - unfolding equation_minus_iff by auto - have **: "\x y::'a. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" - by auto - note scaleR_right_distrib [simp] - show ?t1 - unfolding midpoint_def dist_norm - apply (rule **) - apply (simp add: scaleR_right_diff_distrib) - apply (simp add: scaleR_2) - done - show ?t2 - unfolding midpoint_def dist_norm - apply (rule *) - apply (simp add: scaleR_right_diff_distrib) - apply (simp add: scaleR_2) - done - show ?t3 - unfolding midpoint_def dist_norm - apply (rule *) - apply (simp add: scaleR_right_diff_distrib) - apply (simp add: scaleR_2) - done - show ?t4 - unfolding midpoint_def dist_norm - apply (rule **) - apply (simp add: scaleR_right_diff_distrib) - apply (simp add: scaleR_2) - done -qed - -lemma midpoint_eq_endpoint [simp]: - "midpoint a b = a \ a = b" - "midpoint a b = b \ a = b" - unfolding midpoint_eq_iff by auto - -lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b" - using midpoint_eq_iff by metis - -lemma midpoint_linear_image: - "linear f \ midpoint(f a)(f b) = f(midpoint a b)" -by (simp add: linear_iff midpoint_def) - - -subsection \Line segments\ - -definition\<^marker>\tag important\ closed_segment :: "'a::real_vector \ 'a \ 'a set" - where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" - -definition\<^marker>\tag important\ open_segment :: "'a::real_vector \ 'a \ 'a set" where - "open_segment a b \ closed_segment a b - {a,b}" - -lemmas segment = open_segment_def closed_segment_def - -lemma in_segment: - "x \ closed_segment a b \ (\u. 0 \ u \ u \ 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" - "x \ open_segment a b \ a \ b \ (\u. 0 < u \ u < 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" - using less_eq_real_def by (auto simp: segment algebra_simps) - -lemma closed_segment_linear_image: - "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f" -proof - - interpret linear f by fact - show ?thesis - by (force simp add: in_segment add scale) -qed - -lemma open_segment_linear_image: - "\linear f; inj f\ \ open_segment (f a) (f b) = f ` (open_segment a b)" - by (force simp: open_segment_def closed_segment_linear_image inj_on_def) - -lemma closed_segment_translation: - "closed_segment (c + a) (c + b) = image (\x. c + x) (closed_segment a b)" -apply safe -apply (rule_tac x="x-c" in image_eqI) -apply (auto simp: in_segment algebra_simps) -done - -lemma open_segment_translation: - "open_segment (c + a) (c + b) = image (\x. c + x) (open_segment a b)" -by (simp add: open_segment_def closed_segment_translation translation_diff) - -lemma closed_segment_of_real: - "closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y" - apply (auto simp: image_iff in_segment scaleR_conv_of_real) - apply (rule_tac x="(1-u)*x + u*y" in bexI) - apply (auto simp: in_segment) - done - -lemma open_segment_of_real: - "open_segment (of_real x) (of_real y) = of_real ` open_segment x y" - apply (auto simp: image_iff in_segment scaleR_conv_of_real) - apply (rule_tac x="(1-u)*x + u*y" in bexI) - apply (auto simp: in_segment) - done - -lemma closed_segment_Reals: - "\x \ Reals; y \ Reals\ \ closed_segment x y = of_real ` closed_segment (Re x) (Re y)" - by (metis closed_segment_of_real of_real_Re) - -lemma open_segment_Reals: - "\x \ Reals; y \ Reals\ \ open_segment x y = of_real ` open_segment (Re x) (Re y)" - by (metis open_segment_of_real of_real_Re) - -lemma open_segment_PairD: - "(x, x') \ open_segment (a, a') (b, b') - \ (x \ open_segment a b \ a = b) \ (x' \ open_segment a' b' \ a' = b')" - by (auto simp: in_segment) - -lemma closed_segment_PairD: - "(x, x') \ closed_segment (a, a') (b, b') \ x \ closed_segment a b \ x' \ closed_segment a' b'" - by (auto simp: closed_segment_def) - -lemma closed_segment_translation_eq [simp]: - "d + x \ closed_segment (d + a) (d + b) \ x \ closed_segment a b" -proof - - have *: "\d x a b. x \ closed_segment a b \ d + x \ closed_segment (d + a) (d + b)" - apply (simp add: closed_segment_def) - apply (erule ex_forward) - apply (simp add: algebra_simps) - done - show ?thesis - using * [where d = "-d"] * - by (fastforce simp add:) -qed - -lemma open_segment_translation_eq [simp]: - "d + x \ open_segment (d + a) (d + b) \ x \ open_segment a b" - by (simp add: open_segment_def) - -lemma of_real_closed_segment [simp]: - "of_real x \ closed_segment (of_real a) (of_real b) \ x \ closed_segment a b" - apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward) - using of_real_eq_iff by fastforce - -lemma of_real_open_segment [simp]: - "of_real x \ open_segment (of_real a) (of_real b) \ x \ open_segment a b" - apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE) - using of_real_eq_iff by fastforce - -lemma convex_contains_segment: - "convex S \ (\a\S. \b\S. closed_segment a b \ S)" - unfolding convex_alt closed_segment_def by auto - -lemma closed_segment_in_Reals: - "\x \ closed_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" - by (meson subsetD convex_Reals convex_contains_segment) - -lemma open_segment_in_Reals: - "\x \ open_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" - by (metis Diff_iff closed_segment_in_Reals open_segment_def) - -lemma closed_segment_subset: "\x \ S; y \ S; convex S\ \ closed_segment x y \ S" - by (simp add: convex_contains_segment) - -lemma closed_segment_subset_convex_hull: - "\x \ convex hull S; y \ convex hull S\ \ closed_segment x y \ convex hull S" - using convex_contains_segment by blast - -lemma segment_convex_hull: - "closed_segment a b = convex hull {a,b}" -proof - - have *: "\x. {x} \ {}" by auto - show ?thesis - unfolding segment convex_hull_insert[OF *] convex_hull_singleton - by (safe; rule_tac x="1 - u" in exI; force) -qed - -lemma open_closed_segment: "u \ open_segment w z \ u \ closed_segment w z" - by (auto simp add: closed_segment_def open_segment_def) - -lemma segment_open_subset_closed: - "open_segment a b \ closed_segment a b" - by (auto simp: closed_segment_def open_segment_def) - -lemma bounded_closed_segment: - fixes a :: "'a::euclidean_space" shows "bounded (closed_segment a b)" - by (simp add: segment_convex_hull compact_convex_hull compact_imp_bounded) - -lemma bounded_open_segment: - fixes a :: "'a::euclidean_space" shows "bounded (open_segment a b)" - by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed]) - -lemmas bounded_segment = bounded_closed_segment open_closed_segment - -lemma ends_in_segment [iff]: "a \ closed_segment a b" "b \ closed_segment a b" - unfolding segment_convex_hull - by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) - -lemma eventually_closed_segment: - fixes x0::"'a::real_normed_vector" - assumes "open X0" "x0 \ X0" - shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" -proof - - from openE[OF assms] - obtain e where e: "0 < e" "ball x0 e \ X0" . - then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" - by (auto simp: dist_commute eventually_at) - then show ?thesis - proof eventually_elim - case (elim x) - have "x0 \ ball x0 e" using \e > 0\ by simp - from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] - have "closed_segment x0 x \ ball x0 e" . - also note \\ \ X0\ - finally show ?case . - qed -qed - -lemma segment_furthest_le: - fixes a b x y :: "'a::euclidean_space" - assumes "x \ closed_segment a b" - shows "norm (y - x) \ norm (y - a) \ norm (y - x) \ norm (y - b)" -proof - - obtain z where "z \ {a, b}" "norm (x - y) \ norm (z - y)" - using simplex_furthest_le[of "{a, b}" y] - using assms[unfolded segment_convex_hull] - by auto - then show ?thesis - by (auto simp add:norm_minus_commute) -qed - -lemma closed_segment_commute: "closed_segment a b = closed_segment b a" -proof - - have "{a, b} = {b, a}" by auto - thus ?thesis - by (simp add: segment_convex_hull) -qed - -lemma segment_bound1: - assumes "x \ closed_segment a b" - shows "norm (x - a) \ norm (b - a)" -proof - - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" - using assms by (auto simp add: closed_segment_def) - then show "norm (x - a) \ norm (b - a)" - apply clarify - apply (auto simp: algebra_simps) - apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le) - done -qed - -lemma segment_bound: - assumes "x \ closed_segment a b" - shows "norm (x - a) \ norm (b - a)" "norm (x - b) \ norm (b - a)" -apply (simp add: assms segment_bound1) -by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1) - -lemma open_segment_commute: "open_segment a b = open_segment b a" -proof - - have "{a, b} = {b, a}" by auto - thus ?thesis - by (simp add: closed_segment_commute open_segment_def) -qed - -lemma closed_segment_idem [simp]: "closed_segment a a = {a}" - unfolding segment by (auto simp add: algebra_simps) - -lemma open_segment_idem [simp]: "open_segment a a = {}" - by (simp add: open_segment_def) - -lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \ {a,b}" - using open_segment_def by auto - -lemma convex_contains_open_segment: - "convex s \ (\a\s. \b\s. open_segment a b \ s)" - by (simp add: convex_contains_segment closed_segment_eq_open) - -lemma closed_segment_eq_real_ivl: - fixes a b::real - shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" -proof - - have "b \ a \ closed_segment b a = {b .. a}" - and "a \ b \ closed_segment a b = {a .. b}" - by (auto simp: convex_hull_eq_real_cbox segment_convex_hull) - thus ?thesis - by (auto simp: closed_segment_commute) -qed - -lemma open_segment_eq_real_ivl: - fixes a b::real - shows "open_segment a b = (if a \ b then {a<..x. (v - u) * x + u) ` {0..1}" - by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) - -lemma dist_in_closed_segment: - fixes a :: "'a :: euclidean_space" - assumes "x \ closed_segment a b" - shows "dist x a \ dist a b \ dist x b \ dist a b" -proof (intro conjI) - obtain u where u: "0 \ u" "u \ 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" - using assms by (force simp: in_segment algebra_simps) - have "dist x a = u * dist a b" - apply (simp add: dist_norm algebra_simps x) - by (metis \0 \ u\ abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) - also have "... \ dist a b" - by (simp add: mult_left_le_one_le u) - finally show "dist x a \ dist a b" . - have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" - by (simp add: dist_norm algebra_simps x) - also have "... = (1-u) * dist a b" - proof - - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" - using \u \ 1\ by force - then show ?thesis - by (simp add: dist_norm real_vector.scale_right_diff_distrib) - qed - also have "... \ dist a b" - by (simp add: mult_left_le_one_le u) - finally show "dist x b \ dist a b" . -qed - -lemma dist_in_open_segment: - fixes a :: "'a :: euclidean_space" - assumes "x \ open_segment a b" - shows "dist x a < dist a b \ dist x b < dist a b" -proof (intro conjI) - obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" - using assms by (force simp: in_segment algebra_simps) - have "dist x a = u * dist a b" - apply (simp add: dist_norm algebra_simps x) - by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \0 < u\) - also have *: "... < dist a b" - by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \u < 1\) - finally show "dist x a < dist a b" . - have ab_ne0: "dist a b \ 0" - using * by fastforce - have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" - by (simp add: dist_norm algebra_simps x) - also have "... = (1-u) * dist a b" - proof - - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" - using \u < 1\ by force - then show ?thesis - by (simp add: dist_norm real_vector.scale_right_diff_distrib) - qed - also have "... < dist a b" - using ab_ne0 \0 < u\ by simp - finally show "dist x b < dist a b" . -qed - -lemma dist_decreases_open_segment_0: - fixes x :: "'a :: euclidean_space" - assumes "x \ open_segment 0 b" - shows "dist c x < dist c 0 \ dist c x < dist c b" -proof (rule ccontr, clarsimp simp: not_less) - obtain u where u: "0 \ b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" - using assms by (auto simp: in_segment) - have xb: "x \ b < b \ b" - using u x by auto - assume "norm c \ dist c x" - then have "c \ c \ (c - x) \ (c - x)" - by (simp add: dist_norm norm_le) - moreover have "0 < x \ b" - using u x by auto - ultimately have less: "c \ b < x \ b" - by (simp add: x algebra_simps inner_commute u) - assume "dist c b \ dist c x" - then have "(c - b) \ (c - b) \ (c - x) \ (c - x)" - by (simp add: dist_norm norm_le) - then have "(b \ b) * (1 - u*u) \ 2 * (b \ c) * (1-u)" - by (simp add: x algebra_simps inner_commute) - then have "(1+u) * (b \ b) * (1-u) \ 2 * (b \ c) * (1-u)" - by (simp add: algebra_simps) - then have "(1+u) * (b \ b) \ 2 * (b \ c)" - using \u < 1\ by auto - with xb have "c \ b \ x \ b" - by (auto simp: x algebra_simps inner_commute) - with less show False by auto -qed - -proposition dist_decreases_open_segment: - fixes a :: "'a :: euclidean_space" - assumes "x \ open_segment a b" - shows "dist c x < dist c a \ dist c x < dist c b" -proof - - have *: "x - a \ open_segment 0 (b - a)" using assms - by (metis diff_self open_segment_translation_eq uminus_add_conv_diff) - show ?thesis - using dist_decreases_open_segment_0 [OF *, of "c-a"] assms - by (simp add: dist_norm) -qed - -corollary open_segment_furthest_le: - fixes a b x y :: "'a::euclidean_space" - assumes "x \ open_segment a b" - shows "norm (y - x) < norm (y - a) \ norm (y - x) < norm (y - b)" - by (metis assms dist_decreases_open_segment dist_norm) - -corollary dist_decreases_closed_segment: - fixes a :: "'a :: euclidean_space" - assumes "x \ closed_segment a b" - shows "dist c x \ dist c a \ dist c x \ dist c b" -apply (cases "x \ open_segment a b") - using dist_decreases_open_segment less_eq_real_def apply blast -by (metis DiffI assms empty_iff insertE open_segment_def order_refl) - -lemma convex_intermediate_ball: - fixes a :: "'a :: euclidean_space" - shows "\ball a r \ T; T \ cball a r\ \ convex T" -apply (simp add: convex_contains_open_segment, clarify) -by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment) - -lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \ closed_segment a b" - apply (clarsimp simp: midpoint_def in_segment) - apply (rule_tac x="(1 + u) / 2" in exI) - apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib) - by (metis field_sum_of_halves scaleR_left.add) - -lemma notin_segment_midpoint: - fixes a :: "'a :: euclidean_space" - shows "a \ b \ a \ closed_segment (midpoint a b) b" -by (auto simp: dist_midpoint dest!: dist_in_closed_segment) - -lemma segment_to_closest_point: - fixes S :: "'a :: euclidean_space set" - shows "\closed S; S \ {}\ \ open_segment a (closest_point S a) \ S = {}" - apply (subst disjoint_iff_not_equal) - apply (clarify dest!: dist_in_open_segment) - by (metis closest_point_le dist_commute le_less_trans less_irrefl) - -lemma segment_to_point_exists: - fixes S :: "'a :: euclidean_space set" - assumes "closed S" "S \ {}" - obtains b where "b \ S" "open_segment a b \ S = {}" - by (metis assms segment_to_closest_point closest_point_exists that) - -subsubsection\More lemmas, especially for working with the underlying formula\ - -lemma segment_eq_compose: - fixes a :: "'a :: real_vector" - shows "(\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\x. a + x) o (\u. u *\<^sub>R (b - a))" - by (simp add: o_def algebra_simps) - -lemma segment_degen_1: - fixes a :: "'a :: real_vector" - shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \ a=b \ u=1" -proof - - { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b" - then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b" - by (simp add: algebra_simps) - then have "a=b \ u=1" - by simp - } then show ?thesis - by (auto simp: algebra_simps) -qed - -lemma segment_degen_0: - fixes a :: "'a :: real_vector" - shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \ a=b \ u=0" - using segment_degen_1 [of "1-u" b a] - by (auto simp: algebra_simps) - -lemma add_scaleR_degen: - fixes a b ::"'a::real_vector" - assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \ v" - shows "a=b" - by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms) - -lemma closed_segment_image_interval: - "closed_segment a b = (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}" - by (auto simp: set_eq_iff image_iff closed_segment_def) - -lemma open_segment_image_interval: - "open_segment a b = (if a=b then {} else (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})" - by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1) - -lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval - -lemma open_segment_bound1: - assumes "x \ open_segment a b" - shows "norm (x - a) < norm (b - a)" -proof - - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \ b" - using assms by (auto simp add: open_segment_image_interval split: if_split_asm) - then show "norm (x - a) < norm (b - a)" - apply clarify - apply (auto simp: algebra_simps) - apply (simp add: scaleR_diff_right [symmetric]) - done -qed - -lemma compact_segment [simp]: - fixes a :: "'a::real_normed_vector" - shows "compact (closed_segment a b)" - by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros) - -lemma closed_segment [simp]: - fixes a :: "'a::real_normed_vector" - shows "closed (closed_segment a b)" - by (simp add: compact_imp_closed) - -lemma closure_closed_segment [simp]: - fixes a :: "'a::real_normed_vector" - shows "closure(closed_segment a b) = closed_segment a b" - by simp - -lemma open_segment_bound: - assumes "x \ open_segment a b" - shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)" -apply (simp add: assms open_segment_bound1) -by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute) - -lemma closure_open_segment [simp]: - "closure (open_segment a b) = (if a = b then {} else closed_segment a b)" - for a :: "'a::euclidean_space" -proof (cases "a = b") - case True - then show ?thesis - by simp -next - case False - have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" - apply (rule closure_injective_linear_image [symmetric]) - apply (use False in \auto intro!: injI\) - done - then have "closure - ((\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) = - (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}" - using closure_translation [of a "((\x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"] - by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image) - then show ?thesis - by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan) -qed - -lemma closed_open_segment_iff [simp]: - fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \ a = b" - by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2)) - -lemma compact_open_segment_iff [simp]: - fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \ a = b" - by (simp add: bounded_open_segment compact_eq_bounded_closed) - -lemma convex_closed_segment [iff]: "convex (closed_segment a b)" - unfolding segment_convex_hull by(rule convex_convex_hull) - -lemma convex_open_segment [iff]: "convex (open_segment a b)" -proof - - have "convex ((\u. u *\<^sub>R (b - a)) ` {0<..<1})" - by (rule convex_linear_image) auto - then have "convex ((+) a ` (\u. u *\<^sub>R (b - a)) ` {0<..<1})" - by (rule convex_translation) - then show ?thesis - by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right) -qed - -lemmas convex_segment = convex_closed_segment convex_open_segment - -lemma connected_segment [iff]: - fixes x :: "'a :: real_normed_vector" - shows "connected (closed_segment x y)" - by (simp add: convex_connected) - -lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real - by (auto simp: is_interval_convex_1) - -lemma IVT'_closed_segment_real: - fixes f :: "real \ real" - assumes "y \ closed_segment (f a) (f b)" - assumes "continuous_on (closed_segment a b) f" - shows "\x \ closed_segment a b. f x = y" - using IVT'[of f a y b] - IVT'[of "-f" a "-y" b] - IVT'[of f b y a] - IVT'[of "-f" b "-y" a] assms - by (cases "a \ b"; cases "f b \ f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus) - - subsection\Starlike sets\ definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" lemma starlike_UNIV [simp]: "starlike UNIV" by (simp add: starlike_def) lemma convex_imp_starlike: "convex S \ S \ {} \ starlike S" unfolding convex_contains_segment starlike_def by auto lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\\More results about segments\ lemma dist_half_times2: fixes a :: "'a :: real_normed_vector" shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)" proof - have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))" by simp also have "... = norm ((a + b) - 2 *\<^sub>R x)" by (simp add: real_vector.scale_right_diff_distrib) finally show ?thesis by (simp only: dist_norm) qed lemma closed_segment_as_ball: "closed_segment a b = affine hull {a,b} \ cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) \ norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \ norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) \ norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) \ norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ \ 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" by auto finally show ?thesis . qed show ?thesis by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *) qed lemma open_segment_as_ball: "open_segment a b = affine hull {a,b} \ ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) < norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ < 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" by auto finally show ?thesis . qed show ?thesis using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *) qed lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball lemma closed_segment_neq_empty [simp]: "closed_segment a b \ {}" by auto lemma open_segment_eq_empty [simp]: "open_segment a b = {} \ a = b" proof - { assume a1: "open_segment a b = {}" have "{} \ {0::real<..<1}" by simp then have "a = b" using a1 open_segment_image_interval by fastforce } then show ?thesis by auto qed lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \ a = b" using open_segment_eq_empty by blast lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty lemma inj_segment: fixes a :: "'a :: real_vector" assumes "a \ b" shows "inj_on (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I" proof fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)" by (simp add: algebra_simps) with assms show "x = y" by (simp add: real_vector.scale_right_imp_eq) qed lemma finite_closed_segment [simp]: "finite(closed_segment a b) \ a = b" apply auto apply (rule ccontr) apply (simp add: segment_image_interval) using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast done lemma finite_open_segment [simp]: "finite(open_segment a b) \ a = b" by (auto simp: open_segment_def) lemmas finite_segment = finite_closed_segment finite_open_segment lemma closed_segment_eq_sing: "closed_segment a b = {c} \ a = c \ b = c" by auto lemma open_segment_eq_sing: "open_segment a b \ {c}" by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval) lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing lemma subset_closed_segment: "closed_segment a b \ closed_segment c d \ a \ closed_segment c d \ b \ closed_segment c d" by auto (meson contra_subsetD convex_closed_segment convex_contains_segment) lemma subset_co_segment: "closed_segment a b \ open_segment c d \ a \ open_segment c d \ b \ open_segment c d" using closed_segment_subset by blast lemma subset_open_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ open_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" (is "?lhs = ?rhs") proof (cases "a = b") case True then show ?thesis by simp next case False show ?thesis proof assume rhs: ?rhs with \a \ b\ have "c \ d" using closed_segment_idem singleton_iff by auto have "\uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) = (1 - uc) *\<^sub>R c + uc *\<^sub>R d \ 0 < uc \ uc < 1" if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \ (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \ d" and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d" and u: "0 < u" "u < 1" and uab: "0 \ ua" "ua \ 1" "0 \ ub" "ub \ 1" for u ua ub proof - have "ua \ ub" using neq by auto moreover have "(u - 1) * ua \ 0" using u uab by (simp add: mult_nonpos_nonneg) ultimately have lt: "(u - 1) * ua < u * ub" using u uab by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less) have "p * ua + q * ub < p+q" if p: "0 < p" and q: "0 < q" for p q proof - have "\ p \ 0" "\ q \ 0" using p q not_less by blast+ then show ?thesis by (metis \ua \ ub\ add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5) less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4)) qed then have "(1 - u) * ua + u * ub < 1" using u \ua \ ub\ by (metis diff_add_cancel diff_gt_0_iff_gt) with lt show ?thesis by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps) qed with rhs \a \ b\ \c \ d\ show ?lhs unfolding open_segment_image_interval closed_segment_def by (fastforce simp add:) next assume lhs: ?lhs with \a \ b\ have "c \ d" by (meson finite_open_segment rev_finite_subset) have "closure (open_segment a b) \ closure (open_segment c d)" using lhs closure_mono by blast then have "closed_segment a b \ closed_segment c d" by (simp add: \a \ b\ \c \ d\) then show ?rhs by (force simp: \a \ b\) qed qed lemma subset_oc_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" apply (simp add: subset_open_segment [symmetric]) apply (rule iffI) apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment) apply (meson dual_order.trans segment_open_subset_closed) done lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment subsection\Betweenness\ definition\<^marker>\tag important\ "between = (\(a,b) x. x \ closed_segment a b)" lemma betweenI: assumes "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" shows "between (a, b) x" using assms unfolding between_def closed_segment_def by auto lemma betweenE: assumes "between (a, b) x" obtains u where "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms unfolding between_def closed_segment_def by auto lemma between_implies_scaled_diff: assumes "between (S, T) X" "between (S, T) Y" "S \ Y" obtains c where "(X - Y) = c *\<^sub>R (S - Y)" proof - from \between (S, T) X\ obtain u\<^sub>X where X: "X = u\<^sub>X *\<^sub>R S + (1 - u\<^sub>X) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) from \between (S, T) Y\ obtain u\<^sub>Y where Y: "Y = u\<^sub>Y *\<^sub>R S + (1 - u\<^sub>Y) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) have "X - Y = (u\<^sub>X - u\<^sub>Y) *\<^sub>R (S - T)" proof - from X Y have "X - Y = u\<^sub>X *\<^sub>R S - u\<^sub>Y *\<^sub>R S + ((1 - u\<^sub>X) *\<^sub>R T - (1 - u\<^sub>Y) *\<^sub>R T)" by simp also have "\ = (u\<^sub>X - u\<^sub>Y) *\<^sub>R S - (u\<^sub>X - u\<^sub>Y) *\<^sub>R T" by (simp add: scaleR_left.diff) finally show ?thesis by (simp add: real_vector.scale_right_diff_distrib) qed moreover from Y have "S - Y = (1 - u\<^sub>Y) *\<^sub>R (S - T)" by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) moreover note \S \ Y\ ultimately have "(X - Y) = ((u\<^sub>X - u\<^sub>Y) / (1 - u\<^sub>Y)) *\<^sub>R (S - Y)" by auto from this that show thesis by blast qed lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" unfolding between_def by auto lemma between: "between (a, b) (x::'a::euclidean_space) \ dist a b = (dist a x) + (dist x b)" proof (cases "a = b") case True then show ?thesis by (auto simp add: between_def dist_commute) next case False then have Fal: "norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" by auto have *: "\u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps) have "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" if "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" for u proof - have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" unfolding that(1) by (auto simp add:algebra_simps) show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" unfolding norm_minus_commute[of x a] * using \0 \ u\ \u \ 1\ by simp qed moreover have "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" if "dist a b = dist a x + dist x b" proof - let ?\ = "norm (a - x) / norm (a - b)" show "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" proof (intro exI conjI) show "?\ \ 1" using Fal2 unfolding that[unfolded dist_norm] norm_ge_zero by auto show "x = (1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b" proof (subst euclidean_eq_iff; intro ballI) fix i :: 'a assume i: "i \ Basis" have "((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i = ((norm (a - b) - norm (a - x)) * (a \ i) + norm (a - x) * (b \ i)) / norm (a - b)" using Fal by (auto simp add: field_simps inner_simps) also have "\ = x\i" apply (rule divide_eq_imp[OF Fal]) unfolding that[unfolded dist_norm] using that[unfolded dist_triangle_eq] i apply (subst (asm) euclidean_eq_iff) apply (auto simp add: field_simps inner_simps) done finally show "x \ i = ((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i" by auto qed qed (use Fal2 in auto) qed ultimately show ?thesis by (force simp add: between_def closed_segment_def dist_triangle_eq) qed lemma between_midpoint: fixes a :: "'a::euclidean_space" shows "between (a,b) (midpoint a b)" (is ?t1) and "between (b,a) (midpoint a b)" (is ?t2) proof - have *: "\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by auto show ?t1 ?t2 unfolding between midpoint_def dist_norm by (auto simp add: field_simps inner_simps euclidean_eq_iff[where 'a='a] intro!: *) qed lemma between_mem_convex_hull: "between (a,b) x \ x \ convex hull {a,b}" unfolding between_mem_segment segment_convex_hull .. lemma between_triv_iff [simp]: "between (a,a) b \ a=b" by (auto simp: between_def) lemma between_triv1 [simp]: "between (a,b) a" by (auto simp: between_def) lemma between_triv2 [simp]: "between (a,b) b" by (auto simp: between_def) lemma between_commute: "between (a,b) = between (b,a)" by (auto simp: between_def closed_segment_commute) lemma between_antisym: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) b\ \ a = b" by (auto simp: between dist_commute) lemma between_trans: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) d\ \ between (b,c) d" using dist_triangle2 [of b c d] dist_triangle3 [of b d a] by (auto simp: between dist_commute) lemma between_norm: fixes a :: "'a :: euclidean_space" shows "between (a,b) x \ norm(x - a) *\<^sub>R (b - x) = norm(b - x) *\<^sub>R (x - a)" by (auto simp: between dist_triangle_eq norm_minus_commute algebra_simps) lemma between_swap: fixes A B X Y :: "'a::euclidean_space" assumes "between (A, B) X" assumes "between (A, B) Y" shows "between (X, B) Y \ between (A, Y) X" using assms by (auto simp add: between) lemma between_translation [simp]: "between (a + y,a + z) (a + x) \ between (y,z) x" by (auto simp: between_def) lemma between_trans_2: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,b) d\ \ between (c,d) a" by (metis between_commute between_swap between_trans) lemma between_scaleR_lift [simp]: fixes v :: "'a::euclidean_space" shows "between (a *\<^sub>R v, b *\<^sub>R v) (c *\<^sub>R v) \ v = 0 \ between (a, b) c" by (simp add: between dist_norm scaleR_left_diff_distrib [symmetric] distrib_right [symmetric]) lemma between_1: fixes x::real shows "between (a,b) x \ (a \ x \ x \ b) \ (b \ x \ x \ a)" by (auto simp: between_mem_segment closed_segment_eq_real_ivl) subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "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 unfolding norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally show "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) apply auto done qed (insert \e>0\ \d>0\, auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x]) apply (auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "z \ interior S" apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) by simp (simp add: field_simps norm_minus_commute) then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma closure_open_Int_superset: assumes "open S" "S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(S \ T)" by (metis assms closed_closure closure_minimal inf.orderE open_Int_closure_subset) then show ?thesis by (simp add: closure_mono dual_order.antisym) qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" apply (clarsimp simp add: min_def a) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a x]) using \0 < e\ False apply (auto simp: field_split_simps) done qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" apply (rule sum.cong) apply rule apply auto done also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "rel_interior (convex hull (insert 0 ?p)) = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 ?p)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 ?p))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 ?p)" using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto then have as [rule_format]: "\y. dist x y < e \ (\i\Basis. i \ D \ y\i = 0) \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e / 2) *\<^sub>R i) \ j" apply - apply (rule as[THEN conjunct1]) using D \e > 0\ x0 apply (auto simp: dist_norm inner_simps inner_Basis) done then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D unfolding dist_norm by auto have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e / 2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e / 2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e / 2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 ?p)" using as assms unfolding substd_simplex[OF assms] by fastforce obtain a where a: "a \ D" using \D \ {}\ by auto let ?d = "(1 - sum ((\) x) D) / real (card D)" have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp add: Min_gr_iff) moreover have "?d > 0" using as using \0 < card D\ by auto ultimately have h3: "min (Min (((\) x) ` D)) ?d > 0" by auto have "x \ rel_interior (convex hull (insert 0 ?p))" unfolding rel_interior_ball mem_Collect_eq h0 apply (rule,rule h2) unfolding substd_simplex[OF assms] apply (rule_tac x="min (Min (((\) x) ` D)) ?d" in exI) apply (rule, rule h3) apply safe unfolding mem_ball proof - fix y :: 'a assume y: "dist x y < min (Min ((\) x ` D)) ?d" assume y2: "\i\Basis. i \ D \ y\i = 0" have "sum ((\) y) D \ sum (\i. x\i + ?d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < ?d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] using Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (insert y2, auto) qed } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?D = D let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D" have "finite D" apply (rule finite_subset) using assms(2) apply auto done then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = inverse (2 * real (card D))" apply (rule trans[of _ "sum (\j. if i = j then inverse (2 * real (card D)) else 0) ?D"]) unfolding inner_sum_left apply (rule sum.cong) using \i \ D\ \finite D\ sum.delta'[of D i "(\k. inverse (2 * real (card D)))"] d1 assms(2) by (auto simp: inner_Basis rev_subsetD[OF _ assms(2)]) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real (card D))) ?D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) using \bounded_linear f\ fd * apply auto done ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] apply auto apply blast done moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis using rel_interior_empty by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \x \ a\ apply simp done } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto next case True then have "rel_interior S = {}" using rel_interior_empty by auto then have "closure (rel_interior S) = {}" using closure_empty by auto with True show ?thesis by auto qed qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def apply (simp add: algebra_simps) using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] apply auto done finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" using e_def assms by auto moreover have "e \ 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using mem_cball y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 \ 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by auto qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis apply (simp add: rel_interior_eq openin_open) apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI) apply (simp add: open_segment_as_ball) done qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment lemma starlike_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" shows "starlike T" proof - have "rel_interior S \ {}" by (simp add: assms rel_interior_eq_empty) then obtain a where a: "a \ rel_interior S" by blast with ST have "a \ T" by blast have *: "\x. x \ T \ open_segment a x \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ a]) using assms by blast show ?thesis unfolding starlike_def apply (rule bexI [OF _ \a \ T\]) apply (simp add: closed_segment_eq_open) apply (intro conjI ballI a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) apply (simp add: order_trans [OF * ST]) done qed subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis apply simp by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma closed_affine_hull [iff]: fixes S :: "'n::euclidean_space set" shows "closed (affine hull S)" by (metis affine_affine_hull affine_closed) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) unfolding rel_frontier_def using * closed_affine_hull apply auto done qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" apply (rule affine_dim_equal) using * affine_affine_hull apply auto using \S1 \ {}\ hull_subset[of S1] apply auto using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] apply auto done obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" using x1_def x2_def apply (auto simp add: algebra_simps) using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] apply auto done then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp add: aff_dim_UNIV) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma closure_Int: "closure (\I) \ \{closure S |S. S \ I}" proof - { fix y assume "y \ \I" then have y: "\S \ I. y \ S" by auto { fix S assume "S \ I" then have "y \ closure S" using closure_subset y by auto } then have "y \ \{closure S |S. S \ I}" by auto } then have "\I \ \{closure S |S. S \ I}" by auto moreover have "closed (\{closure S |S. S \ I})" unfolding closed_Inter closed_closure by auto ultimately show ?thesis using closure_hull[of "\I"] hull_minimal[of "\I" "\{closure S |S. S \ I}" "closed"] by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" apply (rule_tac x="z" in exI) using \y \ x\ z_def * e1 e dist_norm[of z y] apply simp done } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" apply (rule convex_Inter) using assms convex_rel_interior apply auto done ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto qed lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto qed lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" shows "rel_interior(S \ T) = interior S \ T" proof - obtain a where aS: "a \ interior S" and aT:"a \ T" using assms by force have "rel_interior S = interior S" by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) then show ?thesis by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) qed lemma closure_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "rel_interior S \ T \ {}" shows "closure(S \ T) = closure S \ T" proof have "closure (S \ T) \ closure T" by (simp add: closure_mono) also have "... \ T" by (simp add: affine_closed assms) finally show "closure(S \ T) \ closure S \ T" by (simp add: closure_mono) next obtain a where "a \ rel_interior S" "a \ T" using assms by auto then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" using affine_diffs_subspace rel_interior_subset assms by blast+ show "closure S \ T \ closure (S \ T)" proof fix x assume "x \ closure S \ T" show "x \ closure (S \ T)" proof (cases "x = a") case True then show ?thesis using \a \ S\ \a \ T\ closure_subset by fastforce next case False then have "x \ closure(open_segment a x)" by auto then show ?thesis using \x \ closure S \ T\ assms convex_affine_closure_Int by blast qed qed qed lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_empty rel_interior_eq_empty by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms rel_interior_empty by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof - { assume "S = {}" then have ?thesis by auto } moreover { assume "T = {}" then have ?thesis by auto } moreover { assume "S \ {}" "T \ {}" then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using fst_linear \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using snd_linear \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" apply (subst convex_rel_interior_inter_two[of "S \ (UNIV :: 'm set)" "(UNIV :: 'n set) \ T"]) using * ri assms convex_Times apply auto done finally have ?thesis using * by auto } ultimately show ?thesis by blast qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def using rel_interior_empty by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def using rel_interior_empty by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" apply (rule_tac x="(y, z)" in exI) apply auto done then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] snd_linear conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: rel_interior_empty cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: rel_interior_empty cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \ S)" f c x]) using convex_cone_hull[of "{(1 :: real)} \ S"] conv apply auto done { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] apply (subst convex_sum[of I "convex hull \(S ` I)" c s]) using * assms convex_convex_hull apply auto done qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto then have "x \ ?rhs" apply auto apply (rule_tac x = c in exI) apply (rule_tac x = s in exI) using * c_def s_def p \x \ S i\ apply auto done } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def apply auto done moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def apply (rule cInf_greatest [OF nonMT]) using \ dual_order.strict_implies_order le_less_linear by blast with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" apply (rule subsetD [OF rel_interior_subset inint]) using \l \ 0\ \0 < d\ \0 < \\ by auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" apply (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) using 1 2 \0 < d\ \0 < \\ apply (auto simp: algebra_simps) done qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp then show ?thesis apply (rule ray_to_rel_frontier [OF \bounded S\ _ _ \l \ 0\]) using a affine_hull_nonempty_interior apply blast by (simp add: \interior S = rel_interior S\ frontier_def rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy apply (auto simp: in_segment) apply (rule_tac x="d" in exI) using \0 < d\ that apply (auto simp: algebra_simps) done ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" apply (auto simp: in_segment) apply (rule_tac x="1/d" in exI) apply (auto simp: algebra_simps) done next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ x]) using df rel_frontier_def by auto qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i\I. convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" apply (subst sum_set_cond_linear[of convex]) using rel_interior_sum rel_interior_sing[of "0"] assms apply (auto simp add: convex_set_plus) done lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i\I. convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i { fix i assume "i \ I" then have "s i \ S i" using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto } then have "\i\I. s i \ S i" by auto moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" apply (subst convex_hull_finite_union[of I S]) using assms apply blast using assms apply blast apply rule apply (rule_tac x = "(\i. (1 :: real) / of_nat (card I))" in exI) apply auto done } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def apply auto done moreover have "sum A I = S + T" using A_def I_def unfolding set_plus_def apply auto unfolding set_plus_def apply auto done ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty rel_interior_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) { fix i assume "i \ I" then have "convex (K i)" unfolding K_def apply (subst convex_cone_hull) apply (subst convex_Times) using assms apply auto done } then have convK: "\i\I. convex (K i)" by auto { fix i assume "i \ I" then have "K0 \ K i" unfolding K0_def K_def apply (subst hull_mono) using \\i\I. C0 \ S i\ apply auto done } then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def apply (subst convex_cone_hull) apply (subst convex_Times) unfolding C0_def using convex_convex_hull apply auto done ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" apply (subst cone_convex_hull) using cone_Union[of "K ` I"] apply auto unfolding K_def using cone_cone_hull apply auto done ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" apply (subst convex_hull_finite_union_cones[of I K]) using assms apply blast using False apply blast unfolding K_def apply rule apply (subst convex_cone_hull) apply (subst convex_Times) using assms cone_cone_hull \\i\I. K i \ {}\ K_def apply auto done finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1::real, x) \ rel_interior K0" using K0_def * set_sum_alt[of I "(\i. rel_interior (K i))"] assms k_def cs apply auto apply (rule_tac x = k in exI) apply (simp add: sum_prod) done then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where "0 < e" "ball x e \ interior I" . moreover define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) ultimately have "K \ I" "K < x" "x \ I" using interior_subset[of I] \x \ interior I\ by auto have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "t \ s" shows "convex hull t = affine hull t \ convex hull s" proof - have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto { fix u v x assume uv: "sum u t = 1" "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\v\t. u v *\<^sub>R v)" "x \ t" then have s: "s = (s - t) \ t" \ \split into separate cases\ using assms by auto have [simp]: "(\x\t. v x *\<^sub>R x) + (\x\s - t. v x *\<^sub>R x) = (\x\t. u x *\<^sub>R x)" "sum v t + sum v (s - t) = 1" using uv fin s by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\s. if x \ t then v x - u x else v x) = 0" "(\x\s. (if x \ t then v x - u x else v x) *\<^sub>R x) = 0" using uv fin by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "convex hull t \ affine hull t" using convex_hull_subset_affine_hull by blast moreover have "convex hull t \ convex hull s" using assms hull_mono by blast moreover have "affine hull t \ convex hull s \ convex hull t" using assms apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) apply (drule_tac x=s in spec) apply (auto simp: fin) apply (rule_tac x=u in exI) apply (rename_tac v) apply (drule_tac x="\x. if x \ t then v x - u x else v x" in spec) apply (force)+ done ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "card s = Suc (DIM ('a))" shows "affine hull s = UNIV" proof (cases "s = {}") case True then show ?thesis using assms by simp next case False then obtain a t where t: "a \ t" "s = insert a t" by blast then have fin: "finite t" using assms by (metis finite_insert aff_independent_finite) show ?thesis using assms t fin apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) apply (rule subset_antisym) apply force apply (rule Fun.vimage_subsetD) apply (metis add.commute diff_add_cancel surj_def) apply (rule card_ge_dim_independent) apply (auto simp: card_image inj_on_def dim_subset_UNIV) done qed lemma affine_independent_span_gt: fixes s :: "'a::euclidean_space set" assumes ind: "\ affine_dependent s" and dim: "DIM ('a) < card s" shows "affine hull s = UNIV" apply (rule affine_independent_span_eq [OF ind]) apply (rule antisym) using assms apply auto apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) done lemma empty_interior_affine_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(affine hull s) = {}" using assms apply (induct s rule: finite_induct) apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) apply (rule empty_interior_lowdim) by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) lemma empty_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(convex hull s) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma rel_interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_interior(convex hull s) = {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. s = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ s" "b \ s" "a \ b" then have s: "s = (s - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst s, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull s" "y \ ?rhs" { fix u T a assume ua: "\x\s. 0 \ u x" "sum u s = 1" "\ 0 < u a" "a \ s" and yT: "y = (\x\s. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull s \ {w. \u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = w}" have ua0: "u a = 0" using ua by auto obtain b where b: "b\s" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\s. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\s. u x *\<^sub>R x) \ affine hull s" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull s" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull s" using d e yT by auto then obtain v where "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto then have False using assms apply (simp add: affine_dependent_explicit_finite fs) apply (drule_tac x="\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) using ua b d apply (auto simp: algebra_simps sum_subtractf sum.distrib) done } note * = this have "y \ rel_interior (convex hull s)" using y apply (simp add: mem_rel_interior affine_hull_convex_hull) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes s :: "'a::euclidean_space set" shows "\ affine_dependent s ==> interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) apply (rule trans [of _ "rel_interior(convex hull s)"]) apply (simp add: affine_independent_span_gt rel_interior_interior) by (simp add: rel_interior_convex_hull_explicit) lemma interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card s" and u: "\x. x\s \ 0 < u x" "sum u s = 1" and a: "a \ s" then have cs: "Suc 0 < card s" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ s" "a \ b" proof (cases "s \ {a}") case True then show thesis using cs subset_singletonD by fastforce next case False then show thesis by (blast intro: that) qed have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u s" apply (rule Groups_Big.sum_mono2) using a b u apply (auto simp: less_imp_le aff_independent_finite assms) done finally have "u a < 1" using \b \ s\ u by fastforce } note [simp] = this show ?thesis using assms apply (auto simp: interior_convex_hull_explicit_minimal) apply (rule_tac x=u in exI) apply (auto simp: not_le) done qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" apply (simp add: segment_convex_hull open_segment_def) apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) done next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" apply simp apply (rule affine_independent_span_gt) apply (simp_all add: affine_dependent_def insert_Diff_if) done then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes s :: "'a::euclidean_space set" shows "compact s ==> closure(convex hull s) = convex hull s" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (\x \ s. u x = 0) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) apply (auto simp: convex_hull_finite fs) apply (drule_tac x=u in spec) apply (rule_tac x=u in exI) apply force apply (rename_tac v) apply (rule notE [OF assms]) apply (simp add: affine_dependent_explicit) apply (rule_tac x=s in exI) apply (auto simp: fs) apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done qed lemma frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (DIM ('a) < card s \ (\x \ s. u x = 0)) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card s") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card s \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = \{convex hull (s - {x}) |x. x \ s}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix u a have "\x\s. 0 \ u x \ a \ s \ u a = 0 \ sum u s = 1 \ \x v. x \ s \ (\x\s - {x}. 0 \ v x) \ sum v (s - {x}) = 1 \ (\x\s - {x}. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ s \ \x\s - {a}. 0 \ u x \ sum u (s - {a}) = 1 \ \v. (\x\s. 0 \ v x) \ (\x\s. v x = 0) \ sum v s = 1 \ (\x\s. v x *\<^sub>R x) = (\x\s - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else rel_frontier(convex hull s))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else \{convex hull (s - {x}) |x. x \ s})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ frontier(convex hull s)" proof (cases "affine_dependent s") case True with assms show ?thesis apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) next case False { assume "card s = Suc (card Basis)" then have cs: "Suc 0 < card s" by (simp add: DIM_positive) with subset_singletonD have "\y \ s. y \ x" by (cases "s \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ interior(convex hull s)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes s :: "'a::euclidean_space set" assumes "card s = Suc (DIM ('a))" shows "interior(convex hull s) = {} \ affine_dependent s" proof - { fix a b assume ab: "a \ interior (convex hull s)" "b \ s" "b \ affine hull (s - {b})" then have "interior(affine hull s) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then have False using ab by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) } then show ?thesis using assms apply auto apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) apply (auto simp: affine_dependent_def) done qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar s \ \u v w. s \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear s \ (\u v. s \ affine hull {u,v})" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ s" by auto { fix u assume *: "\x y. \x\s; y\s\ \ \c. x - y = c *\<^sub>R u" have "\u v. s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" apply (rule_tac x=x in exI) apply (rule_tac x="x+u" in exI, clarify) apply (erule exE [OF * [OF x]]) apply (rename_tac c) apply (rule_tac x="1+c" in exI) apply (rule_tac x="-c" in exI) apply (simp add: algebra_simps) done } moreover { fix u v x y assume *: "s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "x\s \ y\s \ \c. x - y = c *\<^sub>R (v-u)" apply (drule subsetD [OF *])+ apply simp apply clarify apply (rename_tac r1 r2) apply (rule_tac x="r1-r2" in exI) apply (simp add: algebra_simps) apply (metis scaleR_left.add) done } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD by auto (metis collinear_2 numeral_2_eq_2) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - have "card s \ 2 \ card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis using assms apply safe apply (simp add: collinear_small collinear_imp_coplanar) apply (safe dest!: card_eq_SucD) apply (auto simp: coplanar_def) apply (metis hull_subset insert_subset) done qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar s" "linear f" shows "coplanar(f ` s)" proof - { fix u v w assume "s \ affine hull {u, v, w}" then have "f ` s \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` s \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: "coplanar s \ coplanar ((\x. a + x) ` s)" unfolding coplanar_def apply clarify apply (rule_tac x="u+a" in exI) apply (rule_tac x="v+a" in exI) apply (rule_tac x="w+a" in exI) using affine_hull_translation [of a "{u,v,w}" for u v w] apply (force simp: add.commute) done lemma coplanar_translation_eq: "coplanar((\x. a + x) ` s) \ coplanar s" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" proof assume "coplanar s" then show "coplanar (f ` s)" unfolding coplanar_def using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms by (meson coplanar_def coplanar_linear_image) next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` s)" then obtain u v w where "f ` s \ affine hull {u, v, w}" by (auto simp: coplanar_def) then have "g ` f ` s \ g ` (affine hull {u, v, w})" by blast then have "s \ g ` (affine hull {u, v, w})" using g by (simp add: Fun.image_comp) then show "coplanar s" unfolding coplanar_def using affine_hull_linear_image [of g "{u,v,w}" for u v w] \linear g\ linear_conv_bounded_linear by fastforce qed (*The HOL Light proof is simply MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; *) lemma coplanar_subset: "\coplanar t; s \ t\ \ coplanar s" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: "\collinear {a,b,c}; a \ b\ \ c \ affine hull {a,b}" unfolding collinear_def apply clarify apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (rename_tac y x) apply (simp add: affine_hull_2) apply (rule_tac x="1 - x/y" in exI) apply (simp add: algebra_simps) done lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" apply (case_tac "a=b", simp) apply (case_tac "a=c") apply (simp add: insert_commute) apply (case_tac "b=c") apply (simp add: insert_commute) apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) apply (metis collinear_3_affine_hull insert_commute)+ done lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a,midpoint a b,b}" apply (auto simp: collinear_3 collinear_lemma) apply (drule_tac x="-1" in spec) apply (simp add: algebra_simps) done lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c} " using collinear_midpoint by blast moreover have "collinear{a,b,c} \ b = midpoint a c \ dist a b = dist b c" apply (auto simp: collinear_3_expand assms dist_midpoint) apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) apply (simp add: algebra_simps) done ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False with assms show ?thesis apply (auto simp: collinear_3 collinear_lemma between_norm) apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) apply (simp add: vector_add_divide_simps eq_vector_fraction_iff real_vector.scale_minus_right [symmetric]) done qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case True then show ?thesis by (auto simp: dist_commute) next case False show ?thesis apply (rule iffI) apply (simp add: between_midpoint(1) dist_midpoint) using False between_imp_collinear midpoint_collinear by blast qed lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ (insert a (insert b S))" for x using that assms collinear_3_expand by fastforce+ show ?lhs unfolding collinear_def apply (rule_tac x="b-a" in exI) apply (clarify dest!: *) by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff) qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_eq_empty [simp]: "affine hull S = {} \ S = {}" using affine_hull_nonempty by blast lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" apply (simp add: affine_hull_2, safe) apply (rule_tac x=v in image_eqI) apply (simp add: algebra_simps) apply (metis scaleR_add_left scaleR_one, simp) apply (rule_tac x="1-u" in exI) apply (simp add: algebra_simps) done lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" shows "\\ collinear{a,b,c}; DIM('a) = 2\ \ interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) apply (rule_tac x="u a" in exI, simp) apply (rule_tac x="u b" in exI, simp) apply (rule_tac x="u c" in exI, simp) apply (rename_tac uu x y z) apply (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) apply simp done subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex apply auto[1] using inner_zero_right by blast lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "UNIV \ {x. a \ x = b} \ a = 0 \ b = 0" apply (drule_tac c = "((b+1) / (a \ a)) *\<^sub>R a" in subsetD) apply simp_all by (metis add_cancel_right_right zero_neq_one) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "{x. a \ x < b} \ {} \ a = 0 \ b \ 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "{x. a \ x \ b} \ {} \ a = 0 \ b < 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf continuous_on_const) show "open {x. f x < 0}" by (simp add: open_Collect_less contf continuous_on_const) show "S \ {x. 0 < f x}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) show "T \ {x. f x < 0}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then show ?thesis apply (rule separation_normal [OF \closed S\]) apply (rule_tac U=U and V=V in that) by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain) using S apply blast by (metis image_iff le_cases nest) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain_gen [of "S k"]) using S apply auto by (meson le_cases nest subsetCE) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case apply (rule_tac x="max m n" in exI, safe) using max.cobounded2 apply blast by (meson le_max_iff_disj) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n apply (rule closed_Int_compact [OF \closed C\]) apply (rule com) using X \K \ S\ \f ` S \ T\ \y \ T\ apply blast apply (rule compact_sequence_with_limit) apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) done then have comf: "compact {a \ K. f a \ insert y (range (\i. f(X(n + i))))}" for n by (simp add: Keq Int_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = {a \ K. f a \ insert y (range (\i. f (X (k + i))))}" apply (rule exE) apply (rule finite_indexed_bound [OF \finite \\ \], assumption, force) done have "X m \ \\" using X le_Suc_ex by (fastforce dest: m) then show ?thesis by blast qed have "\{{a. a \ K \ f a \ insert y (range (\i. f(X(n + i))))} |n. n \ UNIV} \ {}" apply (rule compact_fip_Heine_Borel) using comf apply force using ne apply (simp add: subset_iff del: insert_iff) done then have "\x. x \ (\n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" by blast then show ?thesis apply (simp add: image_iff fX) by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" apply (rule continuous_injective_image_segment_1) apply (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - have "{.. f ` {x,y} \ {}" "{f z<..} \ f ` {x,y} \ {}" using fz fz_notin \x \ S\ \y \ S\ apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm) apply (metis image_eqI less_eq_real_def)+ done then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" apply (rule continuous_injective_image_segment_1) apply (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" apply safe apply (metis (mono_tags, hide_lams) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_inc subsetCE) by (metis (mono_tags, lifting) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis apply (subst *) apply (rule closedin_closed_trans [OF _ closed_UNIV]) apply (rule closedin_compact_projection [OF \compact S\]) by (simp add: clo closedin_closed_Int) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" apply (rule span_mul) apply (rule span_base) apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) apply (auto simp: S T) done with xa image_iff show ?thesis by fastforce qed show "affine hull S \ affine hull (S \ T)" apply (simp add: subset_hull) apply (simp add: \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) apply (force simp: *) done qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp add: aff_dim_UNIV) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force simp: aff_dim_affine_hull) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True show ?thesis using span_zero [of S] apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) apply (auto simp add: \c = 0\) done next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed show ?thesis apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def cong: image_cong_simp, safe) apply (fastforce simp add: inner_distrib intro: xc_im) apply (force simp: intro!: 2) done qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S apply (simp add: hull_redundant cong: aff_dim_affine_hull2) apply (simp add: affine_hull_insert_span_gen hull_inc) by (force simp add: span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert cong: image_cong_simp) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms show ?thesis apply safe using affine_diffs_subspace [OF assms \b \ S\] apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation image_empty image_insert translation_invert) apply force done qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" apply safe using affine_bounded_eq_trivial assms apply fastforce by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq aff_dim_hyperplane) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have *: "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" apply (clarsimp simp add: field_split_simps) using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) show ?thesis apply (simp add: C k_def) using ass aa Int_iff empty_iff by blast qed qed have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" apply (rule compact_imp_fip) apply (blast intro: compact_cball) using closed_halfspace_ge k_def apply blast apply (metis *) done then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) show ?thesis apply (rule_tac a=a and b = "a \ z" in that, simp_all) using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen apply blast apply (simp_all add: \a \ 0\ szx) done qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) by (metis \0 < e\ add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) proof - have [simp]: "finite s" "finite t" using aff_independent_finite assms by blast+ have "sum u (s \ t) = 1 \ (\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" if [simp]: "sum u s = 1" "sum v t = 1" and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x have "sum f (s \ t) = 0" apply (simp add: f_def sum_Un sum_subtractf) apply (simp add: sum.inter_restrict [symmetric] Int_commute) done moreover have "(\x\(s \ t). f x *\<^sub>R x) = 0" apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq cong del: if_weak_cong) done ultimately have "\v. v \ s \ t \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ s \ u x = (if x \ t then v x else 0)" by (simp add: f_def) presburger have "sum u (s \ t) = sum u s" by (simp add: sum.inter_restrict) then have "sum u (s \ t) = 1" using that by linarith moreover have "(\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ affine hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: affine_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: convex_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ fixes s :: "'a::euclidean_space set set" assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") proof - have "finite s" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction s rule: finite_induct) case empty then show ?case by auto next case (insert t F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (t \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff that \finite S\ sum_dd0 unfolding affine_dependent_explicit apply (simp only: not_ex) apply (drule_tac x=S in spec) apply (drule_tac x=dd in spec, simp) done consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" apply (rule subset_antisym) using in_convex_hull_exchange_unique assms apply blast by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp add: affine_independent_2) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" apply (rule collinear_3_trans [OF _ _ \b \ d\]) using d by (auto simp: insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes s :: "'a::euclidean_space set" obtains f where "finite f" "of_nat (card f) + aff_dim s = DIM('a)" "affine hull s = \f" "\h. h \ f \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ s" and indb: "\ affine_dependent b" and eq: "affine hull s = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card1: "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" apply (rule card_image [OF inj_onI]) by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) have card2: "(card (c - b)) + aff_dim s = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim s" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim s" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent aff_dim_UNIV card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis apply (rule_tac f="{}" in that) using True affc apply (simp_all add: eq [symmetric]) by (metis aff_dim_UNIV aff_dim_affine_hull) next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have affeq: "affine hull s = (\x\(\a. c - {a}) ` (c - b). affine hull x)" using \b \ c\ False apply (subst affine_hull_Inter [OF ind, symmetric]) apply (simp add: eq double_diff) done have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed show ?thesis apply (rule_tac f = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" in that) using \finite c\ apply blast apply (simp add: imeq card1 card2) apply (simp add: affeq, clarify) apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) done qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" apply (rule span_mono) using \0 \ S\ add.left_neutral by force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" apply (rule dim_eq_full [THEN iffD1]) apply (rule antisym [OF dim_subset_UNIV]) using DIM_lt apply simp done ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) let ?U = "(+) (c+c) ` {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c + c) ` {x + v |x v. x \ (+) (- c) ` S \ a \ v = 0}" if "u \ S" "b = a \ v" for u v apply (rule_tac x="u+v-c-c" in image_eqI) apply (simp_all add: algebra_simps) apply (rule_tac x="u-c" in exI) apply (rule_tac x="v-c" in exI) apply (simp add: algebra_simps that c) done moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = ?U" by (fastforce simp: algebra_simps) also have "... = range ((+) (c + c))" by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ apply (rule linear_independent_injective_image) using \0 < e\ inj_on_def by fastforce qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" apply (rule dense_complement_openin_affine_hull) apply (simp add: assms rel_interior_aff_dim) using \convex S\ rel_interior_rel_open rel_open by blast then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then have Sclo: "\w. w \ S \ (w + a') \ S" by fastforce show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis apply (rule_tac a' = "a'" and b' = "a' \ z" in that) apply (auto simp: a ba'' inner_left_distrib False Sclo) done qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "{x - a |x. x \ affine hull S} = span {x - a |x. x \ S}" proof - have *: "((\x. x - a) ` (S - {a})) = {x. x + a \ S} - {0}" by (auto simp: algebra_simps) show ?thesis apply (simp add: affine_hull_span2 [OF assms] *) apply (auto simp: algebra_simps) done qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim {x - a |x. x \ S}" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "{x - c |x. x \ S} = {x - a |x. x \ S}" apply safe apply (simp_all only: xy) using mem_affine_3_minus [OF \affine S\] \a \ S\ \c \ S\ apply blast+ done have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim {x - c |x. x \ B}" by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim {x - c | x. x \ affine hull B}" by (simp add: diffs_affine_hull_span \c \ B\ dim_span) also have "... = dim {x - a |x. x \ S}" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` {x - a| x. x \ T}" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` {x - a| x. x \ T}))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim {x - a| x. x \ T})" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis apply (rule_tac x = T in exI) apply (rule_tac x = "ball x e" in exI) using \T \ \\ apply (simp add: closure_minimal) using closed_cball closure_minimal by blast qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed have 1: "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed have 3: "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast show ?thesis apply (rule_tac x="G (a n)" in exI) apply (intro conjI oG n *) using \K \ S\ \range a = K\ apply blast done qed show ?thesis apply (rule that [OF 1 _ 3]) using Fin \K \ S\ \range a = K\ apply (auto simp: odif) done qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x using D2 [OF that] apply clarify apply (rule_tac x="U \ V" in exI) apply (auto intro: that finite_subset [OF *]) done qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) =(S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis apply (subst eq) apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) by auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" apply (rule closedin_closed_trans) apply (rule continuous_closed_graph_gen [OF contf subset_UNIV]) by (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "open_segment a c \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ open_segment a c" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (open_segment a c))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis using Un_closed_segment [OF b] apply (simp add: closed_segment_eq_open) apply (rule equalityI) using assms apply (simp add: b subset_open_segment) using * by (simp add: insert_commute) qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" apply (intro UN_mono order_refl) apply (simp add: cball_subset_ball_iff field_split_simps) done finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" apply (auto simp: orthogonal_comp_def orthogonal_def) apply (simp add: adjoint_works assms(1) inner_commute) by (metis adjoint_works all_zero_iff assms(1) inner_commute) subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end