# HG changeset patch # User paulson # Date 1500555929 -3600 # Node ID 2562f151541c0177c55462d48e0cd87aca1aa6ee # Parent e5995950b98ae220d401e3dde1dff663eccab030 Divided Convex_Euclidean_Space.thy in half, creating new theory Starlike diff -r e5995950b98a -r 2562f151541c src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy Wed Jul 19 22:56:16 2017 +0100 +++ b/src/HOL/Analysis/Continuous_Extension.thy Thu Jul 20 14:05:29 2017 +0100 @@ -5,7 +5,7 @@ section \Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\ theory Continuous_Extension -imports Convex_Euclidean_Space +imports Starlike begin subsection\Partitions of unity subordinate to locally finite open coverings\ diff -r e5995950b98a -r 2562f151541c src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Wed Jul 19 22:56:16 2017 +0100 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Thu Jul 20 14:05:29 2017 +0100 @@ -7242,7886 +7242,4 @@ using \d > 0\ by auto qed - -subsection \Line segments, Starlike Sets, etc.\ - -definition midpoint :: "'a::real_vector \ 'a \ 'a" - where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" - -definition 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 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: - "linear f \ closed_segment (f a) (f b) = f ` (closed_segment a b)" - by (force simp add: in_segment linear_add_cmul) - -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 midpoint_idem [simp]: "midpoint x x = x" - unfolding midpoint_def - unfolding scaleR_right_distrib - unfolding scaleR_left_distrib[symmetric] - by auto - -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\Starlike sets\ - -definition "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" - -lemma starlike_UNIV [simp]: "starlike UNIV" - by (simp add: starlike_def) - -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_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 convex_imp_starlike: - "convex S \ S \ {} \ starlike S" - unfolding convex_contains_segment starlike_def by auto - -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 real_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]: - fixes a :: "'a::euclidean_space" - shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)" -proof - - have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \ b" - apply (rule closure_injective_linear_image [symmetric]) - apply (simp add:) - using that by (simp add: inj_on_def) - then show ?thesis - by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric] - closure_translation image_comp [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 show ?thesis - apply (simp add: open_segment_image_interval segment_eq_compose) - by (metis image_comp convex_translation) -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 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\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 "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 - unfolding between_def split_conv - by (auto simp add: 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) - show ?thesis - unfolding between_def split_conv closed_segment_def mem_Collect_eq - apply rule - apply (elim exE conjE) - apply (subst dist_triangle_eq) - proof - - fix u - assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" - then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" - unfolding as(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 as(2,3) - by (auto simp add: field_simps) - next - assume as: "dist a b = dist a x + dist x b" - have "norm (a - x) / norm (a - b) \ 1" - using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto - then show "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" - apply (rule_tac x="dist a x / dist a b" in exI) - unfolding dist_norm - apply (subst euclidean_eq_iff) - apply rule - defer - apply rule - prefer 3 - apply rule - proof - - fix i :: 'a - assume i: "i \ Basis" - have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^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 as[unfolded dist_norm] - using as[unfolded dist_triangle_eq] - apply - - apply (subst (asm) euclidean_eq_iff) - using i - apply (erule_tac x=i in ballE) - apply (auto simp add: field_simps inner_simps) - done - finally show "x \ i = - ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \ i" - by auto - qed (insert Fal2, auto) - qed -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 - apply(rule_tac[!] *) - unfolding euclidean_eq_iff[where 'a='a] - apply (auto simp add: field_simps inner_simps) - done -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 \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 - apply (rule_tac x="e*d" in exI) - apply rule - defer - unfolding subset_eq Ball_def mem_ball - proof (rule, rule) - fix y - assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" - 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) - apply (auto simp add:field_simps norm_minus_commute) - done - then show ?thesis - unfolding * - apply - - apply (rule mem_interior_convex_shrink) - using assms(1,4-5) \y\s\ - apply auto - done -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: divide_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 \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)}" - unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] - apply (rule set_eqI, rule) - unfolding mem_Collect_eq - apply (erule_tac[!] exE) - apply (erule_tac[!] conjE)+ - unfolding sum_clauses(2)[OF \finite s\] - apply (rule_tac x=u in exI) - defer - apply (rule_tac x="\x. if x = 0 then 1 - sum u s else u x" in exI) - using assms(2) - unfolding if_smult and sum_delta_notmem[OF assms(2)] - apply auto - done - -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\] - apply (rule set_eqI) - unfolding mem_Collect_eq - apply rule - apply (elim exE conjE) - apply (erule_tac[2] conjE)+ - proof - - fix x :: "'a::euclidean_space" - fix u - assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" "(\x\?D. u x *\<^sub>R x) = x" - have *: "\i\Basis. i:d \ u i = x\i" - and "(\i\Basis. i \ d \ x \ i = 0)" - using as(3) - unfolding substdbasis_expansion_unique[OF assms] - by auto - then have **: "sum u ?D = sum (op \ x) ?D" - apply - - apply (rule sum.cong) - using assms - apply auto - done - have "(\i\Basis. 0 \ x\i) \ sum (op \ x) ?D \ 1" - proof (rule,rule) - fix i :: 'a - assume i: "i \ Basis" - have "i \ d \ 0 \ x\i" - unfolding *[rule_format,OF i,symmetric] - apply (rule_tac as(1)[rule_format]) - apply auto - done - moreover have "i \ d \ 0 \ x\i" - using \(\i\Basis. i \ d \ x \ i = 0)\[rule_format, OF i] by auto - ultimately show "0 \ x\i" by auto - qed (insert as(2)[unfolded **], auto) - then show "(\i\Basis. 0 \ x\i) \ sum (op \ x) ?D \ 1 \ (\i\Basis. i \ d \ x \ i = 0)" - using \(\i\Basis. i \ d \ x \ i = 0)\ by auto - next - fix x :: "'a::euclidean_space" - assume as: "\i\Basis. 0 \ x \ i" "sum (op \ x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" - show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" - using as d - unfolding substdbasis_expansion_unique[OF assms] - apply (rule_tac x="inner x" in exI) - apply auto - done - 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}" - apply (rule set_eqI) - unfolding mem_interior std_simplex - unfolding subset_eq mem_Collect_eq Ball_def mem_ball - unfolding Ball_def[symmetric] - apply rule - apply (elim exE conjE) - defer - apply (erule conjE) -proof - - fix x :: 'a - fix e - assume "e > 0" and as: "\xa. dist x xa < e \ (\x\Basis. 0 \ xa \ x) \ sum (op \ xa) Basis \ 1" - show "(\xa\Basis. 0 < x \ xa) \ sum (op \ x) Basis < 1" - apply safe - proof - - fix i :: 'a - assume i: "i \ Basis" - then show "0 < x \ i" - using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \e > 0\ - unfolding dist_norm - by (auto elim!: ballE[where x=i] simp: 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 (op \ (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" - apply (rule_tac sum.cong) - apply auto - done - have "sum (op \ x) Basis < sum (op \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" - unfolding * sum.distrib - using \e > 0\ DIM_positive[where 'a='a] - apply (subst sum.delta') - apply (auto simp: SOME_Basis) - done - also have "\ \ 1" - using ** - apply (drule_tac as[rule_format]) - apply auto - done - finally show "sum (op \ x) Basis < 1" by auto - qed -next - fix x :: 'a - assume as: "\i\Basis. 0 < x \ i" "sum (op \ x) Basis < 1" - obtain a :: 'b where "a \ UNIV" using UNIV_witness .. - let ?d = "(1 - sum (op \ x) Basis) / real (DIM('a))" - show "\e>0. \y. dist x y < e \ (\i\Basis. 0 \ y \ i) \ sum (op \ y) Basis \ 1" - proof (rule_tac x="min (Min ((op \ x) ` Basis)) D" for D in exI, intro conjI impI allI) - fix y - assume y: "dist x y < min (Min (op \ x ` Basis)) ?d" - have "sum (op \ y) Basis \ sum (\i. x\i + ?d) Basis" - proof (rule sum_mono) - fix i :: 'a - assume i: "i \ Basis" - then have "\y\i - x\i\ < ?d" - apply - - apply (rule le_less_trans) - using Basis_le_norm[OF i, of "y - x"] - using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] - apply (auto simp add: norm_minus_commute inner_diff_left) - done - 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 (op \ y) Basis \ 1" . - show "(\i\Basis. 0 \ y \ i)" - proof safe - fix i :: 'a - assume i: "i \ Basis" - have "norm (x - y) < x\i" - apply (rule less_le_trans) - apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) - using i - apply auto - done - 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 ((op \ x) ` Basis) > 0" - using as by simp - moreover have "?d > 0" - using as by (auto simp: Suc_le_eq) - ultimately show "0 < min (Min (op \ x ` Basis)) ((1 - sum (op \ 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 (op \ ?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 (op \ ?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" - apply (rule finite_subset) - using assms - apply auto - done - 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 e0: "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: "\xa. dist x xa < e \ (\i\Basis. i \ d \ xa\i = 0) \ - (\i\d. 0 \ xa \ i) \ sum (op \ xa) 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 (op \ x) d < 1 \ (\i\Basis. i \ d \ x\i = 0)" - apply rule - apply rule - proof - - fix i :: 'a - assume "i \ d" - then have "\ia\d. 0 \ (x - (e / 2) *\<^sub>R i) \ ia" - apply - - apply (rule as[rule_format,THEN conjunct1]) - unfolding dist_norm - using d \e > 0\ x0 - apply (auto simp: inner_simps inner_Basis) - done - then show "0 < x \ i" - apply (erule_tac x=i in ballE) - using \e > 0\ \i \ d\ d - apply (auto simp: inner_simps inner_Basis) - done - 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 (op \ (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 (op \ x) d < sum (op \ (x + (e / 2) *\<^sub>R a)) d" - unfolding * sum.distrib - using \e > 0\ \a \ d\ - using \finite d\ - by (auto simp add: sum.delta') - also have "\ \ 1" - using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] - by auto - finally show "sum (op \ x) d < 1 \ (\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 (op \ x) d) / real (card d)" - have "0 < card d" using \d \ {}\ \finite d\ - by (simp add: card_gt_0_iff) - have "Min ((op \ 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 ((op \ 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 ((op \ 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 (op \ x ` d)) ?d" - assume y2: "\i\Basis. i \ d \ y\i = 0" - have "sum (op \ 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\ < ?d" - apply (rule le_less_trans) - using Basis_le_norm[OF i, of "y - x"] - using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] - apply (auto simp add: norm_minus_commute inner_simps) - done - 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 (op \ 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 "op \ 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 (op \ 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 set_rev_mp[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 (op \ ?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 (op \ ?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_superset[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 \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 auto - 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_inc 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 \ op + (-a) ` S" - using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto - then have "rel_interior (op + (-a) ` S) \ {}" - using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"] - convex_translation[of S "-a"] assms - by auto - then have "rel_interior S \ {}" - using rel_interior_translation by auto - } - 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" - apply (simp add: hull_inc affine_hull_span_0) - using assms dim_eq_full indep_card_eq_dim_span by fastforce - 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 (subtopology euclidean (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)" - apply auto - using assms - apply simp - done - 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 "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" - apply (simp add: rel_frontier_def) - apply (simp add: rel_interior_eq_closure [symmetric]) - using rel_interior_subset_closure by blast - -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 (subtopology euclidean (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 \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 - 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 \linear f\ by (simp add: linear_add_cmul) - 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 - - 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 (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image) - { - 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 (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image) - 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 "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)" - using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S] - linear_conv_bounded_linear[of "op *\<^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 "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^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 ((op *\<^sub>R c) ` S) \ rel_open ((op *\<^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 \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 \ op *\<^sub>R c ` S = S)" - using cone_iff[of S] assms by auto - then have *: "0 \ ({0} \ rel_interior S)" - and "\c. c > 0 \ op *\<^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 \ ((op *\<^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 = (op *\<^sub>R c ` S)" - using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto - then have "rel_interior (f c) = op *\<^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 - have "z \ ?rhs" - using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ - apply auto - apply (rule_tac x = "fst z" in exI) - apply (rule_tac x = x in exI) - using * - apply auto - done - } - 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 - - -subsection \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 - apply auto - apply (rule_tac x = c in exI) - apply (rule_tac x = s in exI) - using cs - apply auto - done - } - 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 simp add: convex_convex_hull) - } - 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\Explicit formulas for interior and relative interior of convex hull\ - -lemma interior_atLeastAtMost [simp]: - fixes a::real shows "interior {a..b} = {a<.. x < b \ (at x within {a..b}) = at x" - by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost) - -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) - apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) - apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card) - done - -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_hull_convex_hull 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\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 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\The infimum of the distance between two sets\ - -definition setdist :: "'a::metric_space set \ 'a set \ real" where - "setdist s t \ - (if s = {} \ t = {} then 0 - else Inf {dist x y| x y. x \ s \ y \ t})" - -lemma setdist_empty1 [simp]: "setdist {} t = 0" - by (simp add: setdist_def) - -lemma setdist_empty2 [simp]: "setdist t {} = 0" - by (simp add: setdist_def) - -lemma setdist_pos_le [simp]: "0 \ setdist s t" - by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) - -lemma le_setdistI: - assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" - shows "d \ setdist s t" - using assms - by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) - -lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" - unfolding setdist_def - by (auto intro!: bdd_belowI [where m=0] cInf_lower) - -lemma le_setdist_iff: - "d \ setdist s t \ - (\x \ s. \y \ t. d \ dist x y) \ (s = {} \ t = {} \ d \ 0)" - apply (cases "s = {} \ t = {}") - apply (force simp add: setdist_def) - apply (intro iffI conjI) - using setdist_le_dist apply fastforce - apply (auto simp: intro: le_setdistI) - done - -lemma setdist_ltE: - assumes "setdist s t < b" "s \ {}" "t \ {}" - obtains x y where "x \ s" "y \ t" "dist x y < b" -using assms -by (auto simp: not_le [symmetric] le_setdist_iff) - -lemma setdist_refl: "setdist s s = 0" - apply (cases "s = {}") - apply (force simp add: setdist_def) - apply (rule antisym [OF _ setdist_pos_le]) - apply (metis all_not_in_conv dist_self setdist_le_dist) - done - -lemma setdist_sym: "setdist s t = setdist t s" - by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) - -lemma setdist_triangle: "setdist s t \ setdist s {a} + setdist {a} t" -proof (cases "s = {} \ t = {}") - case True then show ?thesis - using setdist_pos_le by fastforce -next - case False - have "\x. x \ s \ setdist s t - dist x a \ setdist {a} t" - apply (rule le_setdistI, blast) - using False apply (fastforce intro: le_setdistI) - apply (simp add: algebra_simps) - apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) - done - then have "setdist s t - setdist {a} t \ setdist s {a}" - using False by (fastforce intro: le_setdistI) - then show ?thesis - by (simp add: algebra_simps) -qed - -lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" - by (simp add: setdist_def) - -lemma setdist_Lipschitz: "\setdist {x} s - setdist {y} s\ \ dist x y" - apply (subst setdist_singletons [symmetric]) - by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) - -lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} s))" - by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) - -lemma continuous_on_setdist [continuous_intros]: "continuous_on t (\y. (setdist {y} s))" - by (metis continuous_at_setdist continuous_at_imp_continuous_on) - -lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\y. (setdist {y} s))" - by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) - -lemma setdist_subset_right: "\t \ {}; t \ u\ \ setdist s u \ setdist s t" - apply (cases "s = {} \ u = {}", force) - apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) - done - -lemma setdist_subset_left: "\s \ {}; s \ t\ \ setdist t u \ setdist s u" - by (metis setdist_subset_right setdist_sym) - -lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t" -proof (cases "s = {} \ t = {}") - case True then show ?thesis by force -next - case False - { fix y - assume "y \ t" - have "continuous_on (closure s) (\a. dist a y)" - by (auto simp: continuous_intros dist_norm) - then have *: "\x. x \ closure s \ setdist s t \ dist x y" - apply (rule continuous_ge_on_closure) - apply assumption - apply (blast intro: setdist_le_dist \y \ t\ ) - done - } note * = this - show ?thesis - apply (rule antisym) - using False closure_subset apply (blast intro: setdist_subset_left) - using False * - apply (force simp add: closure_eq_empty intro!: le_setdistI) - done -qed - -lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s" -by (metis setdist_closure_1 setdist_sym) - -lemma setdist_compact_closed: - fixes S :: "'a::euclidean_space set" - assumes S: "compact S" and T: "closed T" - and "S \ {}" "T \ {}" - shows "\x \ S. \y \ T. dist x y = setdist S T" -proof - - have "(\x\ S. \y \ T. {x - y}) \ {}" - using assms by blast - then have "\x \ S. \y \ T. dist x y \ setdist S T" - apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF S T]]) - apply (simp add: dist_norm le_setdist_iff) - apply blast - done - then show ?thesis - by (blast intro!: antisym [OF _ setdist_le_dist] ) -qed - -lemma setdist_closed_compact: - fixes S :: "'a::euclidean_space set" - assumes S: "closed S" and T: "compact T" - and "S \ {}" "T \ {}" - shows "\x \ S. \y \ T. dist x y = setdist S T" - using setdist_compact_closed [OF T S \T \ {}\ \S \ {}\] - by (metis dist_commute setdist_sym) - -lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" - by (metis antisym dist_self setdist_le_dist setdist_pos_le) - -lemma setdist_eq_0_compact_closed: - fixes S :: "'a::euclidean_space set" - assumes S: "compact S" and T: "closed T" - shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" - apply (cases "S = {} \ T = {}", force) - using setdist_compact_closed [OF S T] - apply (force intro: setdist_eq_0I ) - done - -corollary setdist_gt_0_compact_closed: - fixes S :: "'a::euclidean_space set" - assumes S: "compact S" and T: "closed T" - shows "setdist S T > 0 \ (S \ {} \ T \ {} \ S \ T = {})" - using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] - by linarith - -lemma setdist_eq_0_closed_compact: - fixes S :: "'a::euclidean_space set" - assumes S: "closed S" and T: "compact T" - shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" - using setdist_eq_0_compact_closed [OF T S] - by (metis Int_commute setdist_sym) - -lemma setdist_eq_0_bounded: - fixes S :: "'a::euclidean_space set" - assumes "bounded S \ bounded T" - shows "setdist S T = 0 \ S = {} \ T = {} \ closure S \ closure T \ {}" - apply (cases "S = {} \ T = {}", force) - using setdist_eq_0_compact_closed [of "closure S" "closure T"] - setdist_eq_0_closed_compact [of "closure S" "closure T"] assms - apply (force simp add: bounded_closure compact_eq_bounded_closed) - done - -lemma setdist_unique: - "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ - \ setdist S T = dist a b" - by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) - -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 setdist_eq_0_sing_1: - fixes S :: "'a::euclidean_space set" - shows "setdist {x} S = 0 \ S = {} \ x \ closure S" - by (auto simp: setdist_eq_0_bounded) - -lemma setdist_eq_0_sing_2: - fixes S :: "'a::euclidean_space set" - shows "setdist S {x} = 0 \ S = {} \ x \ closure S" - by (auto simp: setdist_eq_0_bounded) - -lemma setdist_neq_0_sing_1: - fixes S :: "'a::euclidean_space set" - shows "\setdist {x} S = a; a \ 0\ \ S \ {} \ x \ closure S" - by (auto simp: setdist_eq_0_sing_1) - -lemma setdist_neq_0_sing_2: - fixes S :: "'a::euclidean_space set" - shows "\setdist S {x} = a; a \ 0\ \ S \ {} \ x \ closure S" - by (auto simp: setdist_eq_0_sing_2) - -lemma setdist_sing_in_set: - fixes S :: "'a::euclidean_space set" - shows "x \ S \ setdist {x} S = 0" - using closure_subset by (auto simp: setdist_eq_0_sing_1) - -lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" - using setdist_subset_left by auto - -lemma setdist_eq_0_closed: - fixes S :: "'a::euclidean_space set" - shows "closed S \ (setdist {x} S = 0 \ S = {} \ x \ S)" -by (simp add: setdist_eq_0_sing_1) - -lemma setdist_eq_0_closedin: - fixes S :: "'a::euclidean_space set" - shows "\closedin (subtopology euclidean u) S; x \ u\ - \ (setdist {x} S = 0 \ S = {} \ x \ S)" - by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) - -subsection\Basic lemmas about hyperplanes and halfspaces\ - -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\Use set distance for an easy proof of separation properties\ - -proposition 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 (subtopology euclidean U) S" - and UT: "closedin (subtopology euclidean U) T" - and "S \ T = {}" - obtains S' T' where "openin (subtopology euclidean U) S'" - "openin (subtopology euclidean U) T'" - "S \ S'" "T \ T'" "S' \ T' = {}" -proof (cases "S = {} \ T = {}") - case True with that show ?thesis - apply safe - using UT closedin_subset apply blast - using US closedin_subset apply blast - done -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' = "{x\U. f x > 0}" and T' = "{x\U. f x < 0}" in that) - show "{x \ U. 0 < f x} \ {x \ U. f x < 0} = {}" - by auto - have "openin (subtopology euclidean U) {x \ U. f x \ {0<..}}" - by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) - then show "openin (subtopology euclidean U) {x \ U. 0 < f x}" by simp - next - have "openin (subtopology euclidean U) {x \ U. f x \ {..<0}}" - by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) - then show "openin (subtopology euclidean U) {x \ U. f x < 0}" by simp - next - show "S \ {x \ U. 0 < f x}" - apply (clarsimp simp add: f_def setdist_sing_in_set) - using assms - by (metis False Int_insert_right closedin_imp_subset empty_not_insert insert_absorb insert_subset linorder_neqE_linordered_idom not_le setdist_eq_0_closedin setdist_pos_le) - show "T \ {x \ U. f x < 0}" - apply (clarsimp simp add: f_def setdist_sing_in_set) - using assms - by (metis False closedin_subset disjoint_iff_not_equal insert_absorb insert_subset linorder_neqE_linordered_idom not_less setdist_eq_0_closedin setdist_pos_le topspace_euclidean_subtopology) - 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\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 (subtopology euclidean S) K" - and com: "\U. \U \ T; compact U\ \ compact {x \ S. f x \ U}" - and "f ` S \ T" - shows "closedin (subtopology euclidean 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 \ {a \ S. f a \ 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 *: "{x \ S. f x \ 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 {x \ f ` S. g x \ C}" if C: "C \ S" "compact C" for C - proof - - obtain h :: "'a set \ 'a" where "h C \ C \ h C \ S \ compact (f ` C)" - by (force simp: C RHS) - moreover have "f ` C = {b \ f ` S. g b \ C}" - using C gf by auto - ultimately show "compact {b \ f ` S. g b \ C}" - using C by auto - qed - show ?lhs - using proper_map [OF _ _ gfim] ** - by (simp add: continuous_on_closed * closedin_imp_subset) -qed - -subsection\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 (subtopology euclidean T) K" - shows "compact {x. x \ S \ f x \ 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 {z \ S \ T. fst z \ K}" -proof - - have "{z \ S \ T. fst z \ 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 (subtopology euclidean (S \ T)) c" - shows "closedin (subtopology euclidean 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 {z \ S \ T. snd z \ K}" -proof - - have "{z \ S \ T. snd z \ 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 (subtopology euclidean (S \ T)) c" - shows "closedin (subtopology euclidean 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 (subtopology euclidean (S \ T)) U" - shows "closedin (subtopology euclidean 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 (subtopology euclidean 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\Representing affine hull as a finite intersection of hyperplanes\ - -proposition 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 \ op + 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 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_superset) - 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 (subtopology euclidean (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 (subtopology euclidean (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 (subtopology euclidean (affine hull S)) S" and "finite F" "F \ S" - shows "affine hull (S - F) = affine hull S" -proof - - have clo: "closedin (subtopology euclidean 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 - -subsection\Properties of special hyperplanes\ - -lemma subspace_hyperplane: "subspace {x. a \ x = 0}" - by (simp add: subspace_def inner_right_distrib) - -lemma subspace_hyperplane2: "subspace {x. x \ a = 0}" - by (simp add: inner_commute inner_right_distrib subspace_def) - -lemma special_hyperplane_span: - fixes S :: "'n::euclidean_space set" - assumes "k \ Basis" - shows "{x. k \ x = 0} = span (Basis - {k})" -proof - - have *: "x \ span (Basis - {k})" if "k \ x = 0" for x - proof - - have "x = (\b\Basis. (x \ b) *\<^sub>R b)" - by (simp add: euclidean_representation) - also have "... = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" - by (auto simp: sum.remove [of _ k] inner_commute assms that) - finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . - then show ?thesis - by (simp add: Linear_Algebra.span_finite) metis - qed - show ?thesis - apply (rule span_subspace [symmetric]) - using assms - apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) - done -qed - -lemma dim_special_hyperplane: - fixes k :: "'n::euclidean_space" - shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" -apply (simp add: special_hyperplane_span) -apply (rule Linear_Algebra.dim_unique [OF subset_refl]) -apply (auto simp: Diff_subset independent_substdbasis) -apply (metis member_remove remove_def span_clauses(1)) -done - -proposition dim_hyperplane: - fixes a :: "'a::euclidean_space" - assumes "a \ 0" - shows "dim {x. a \ x = 0} = DIM('a) - 1" -proof - - have span0: "span {x. a \ x = 0} = {x. a \ x = 0}" - by (rule span_unique) (auto simp: subspace_hyperplane) - then obtain B where "independent B" - and Bsub: "B \ {x. a \ x = 0}" - and subspB: "{x. a \ x = 0} \ span B" - and card0: "(card B = dim {x. a \ x = 0})" - and ortho: "pairwise orthogonal B" - using orthogonal_basis_exists by metis - with assms have "a \ span B" - by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace) - then have ind: "independent (insert a B)" - by (simp add: \independent B\ independent_insert) - have "finite B" - using \independent B\ independent_bound by blast - have "UNIV \ span (insert a B)" - proof fix y::'a - obtain r z where z: "y = r *\<^sub>R a + z" "a \ z = 0" - apply (rule_tac r="(a \ y) / (a \ a)" and z = "y - ((a \ y) / (a \ a)) *\<^sub>R a" in that) - using assms - by (auto simp: algebra_simps) - show "y \ span (insert a B)" - by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq - add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) - qed - then have dima: "DIM('a) = dim(insert a B)" - by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim) - then show ?thesis - by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) -qed - -lemma lowdim_eq_hyperplane: - fixes S :: "'a::euclidean_space set" - assumes "dim S = DIM('a) - 1" - obtains a where "a \ 0" and "span S = {x. a \ x = 0}" -proof - - have [simp]: "dim S < DIM('a)" - by (simp add: DIM_positive assms) - then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" - using lowdim_subset_hyperplane [of S] by fastforce - show ?thesis - using b that real_vector_class.subspace_span [of S] - by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane) -qed - -lemma dim_eq_hyperplane: - fixes S :: "'n::euclidean_space set" - shows "dim S = DIM('n) - 1 \ (\a. a \ 0 \ span S = {x. a \ x = 0})" -by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) - -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 - 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 (rule ex_cong) - apply (metis (mono_tags) span_0 \c = 0\ image_add_0 inner_zero_right mem_Collect_eq) - done - next - case False - have xc_im: "x \ op + 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 \ op + c ` {y. a \ y = 0}" - by blast - qed - have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" - if "op + 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 "op + 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, 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\Some stepping theorems\ - -lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0" - by (force intro!: dim_unique) - -lemma dim_insert: - fixes x :: "'a::euclidean_space" - shows "dim (insert x S) = (if x \ span S then dim S else dim S + 1)" -proof - - show ?thesis - proof (cases "x \ span S") - case True then show ?thesis - by (metis dim_span span_redundant) - next - case False - obtain B where B: "B \ span S" "independent B" "span S \ span B" "card B = dim (span S)" - using basis_exists [of "span S"] by blast - have 1: "insert x B \ span (insert x S)" - by (meson \B \ span S\ dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI) - have 2: "span (insert x S) \ span (insert x B)" - by (metis \B \ span S\ \span S \ span B\ span_breakdown_eq span_subspace subsetI subspace_span) - have 3: "independent (insert x B)" - by (metis B independent_insert span_subspace subspace_span False) - have "dim (span (insert x S)) = Suc (dim S)" - apply (rule dim_unique [OF 1 2 3]) - by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE) - then show ?thesis - by (simp add: False) - qed -qed - -lemma dim_singleton [simp]: - fixes x :: "'a::euclidean_space" - shows "dim{x} = (if x = 0 then 0 else 1)" -by (simp add: dim_insert) - -lemma dim_eq_0 [simp]: - fixes S :: "'a::euclidean_space set" - shows "dim S = 0 \ S \ {0}" -apply safe -apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le) -by (metis dim_singleton dim_subset le_0_eq) - -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) - apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0) - apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff) - done -qed - -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\General case without assuming closure and getting non-strict separation\ - -proposition 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 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_mul span_superset) - 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_mul span_superset) - 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: divide_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 (op + (- z) ` S)" - by (simp add: \convex S\) - moreover have "op + (- z) ` S \ {}" - by (simp add: \S \ {}\) - moreover have "0 \ op + (- z) ` S" - using zno by auto - ultimately obtain a where "a \ span (op + (- z) ` S)" "a \ 0" - and a: "\x. x \ (op + (- 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 (op + (- z) ` S)\ affine_hull_insert_span_gen apply blast - apply (simp_all add: \a \ 0\ szx) - done -qed - -proposition 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)) (op \ 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\ 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 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 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 - 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 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 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 - -subsection\Misc results about span\ - -lemma eq_span_insert_eq: - assumes "(x - y) \ span S" - shows "span(insert x S) = span(insert y S)" -proof - - have *: "span(insert x S) \ span(insert y S)" if "(x - y) \ span S" for x y - proof - - have 1: "(r *\<^sub>R x - r *\<^sub>R y) \ span S" for r - by (metis real_vector.scale_right_diff_distrib span_mul that) - have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for z k - by (simp add: real_vector.scale_right_diff_distrib) - show ?thesis - apply (clarsimp simp add: span_breakdown_eq) - by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq) - qed - show ?thesis - apply (intro subset_antisym * assms) - using assms subspace_neg subspace_span minus_diff_eq by force -qed - -lemma dim_psubset: - fixes S :: "'a :: euclidean_space set" - shows "span S \ span T \ dim S < dim T" -by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) - - -lemma basis_subspace_exists: - fixes S :: "'a::euclidean_space set" - shows - "subspace S - \ \b. finite b \ b \ S \ - independent b \ span b = S \ card b = dim S" -by (metis span_subspace basis_exists independent_imp_finite) - -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_superset) - 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 subs) -qed - -proposition 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 (op + (- c) ` S)" by blast - then have aff: "affine (op + (- c) ` S)" - by (simp add: subspace_imp_affine) - have 0: "0 \ op + (- c) ` S" - by (simp add: \c \ S\) - obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ op + (- c) ` S" - using assms by auto - then have adc: "a \ (d - c) \ 0" - by (simp add: c inner_diff_right) - let ?U = "op + (c+c) ` {x + y |x y. x \ op + (- c) ` S \ a \ y = 0}" - have "u + v \ op + (c + c) ` {x + v |x v. x \ op + (- 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 "... = op + (c+c) ` UNIV" - by (simp add: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) - also have "... = UNIV" - by (simp add: translation_UNIV) - finally show ?thesis . -qed - -proposition dim_sums_Int: - fixes S :: "'a :: euclidean_space set" - assumes "subspace S" "subspace T" - shows "dim {x + y |x y. x \ S \ y \ T} + dim(S \ T) = dim S + dim T" -proof - - obtain B where B: "B \ S \ T" "S \ T \ span B" - and indB: "independent B" - and cardB: "card B = dim (S \ T)" - using basis_exists by blast - then obtain C D where "B \ C" "C \ S" "independent C" "S \ span C" - and "B \ D" "D \ T" "independent D" "T \ span D" - using maximal_independent_subset_extend - by (metis Int_subset_iff \B \ S \ T\ indB) - then have "finite B" "finite C" "finite D" - by (simp_all add: independent_imp_finite indB independent_bound) - have Beq: "B = C \ D" - apply (rule sym) - apply (rule spanning_subset_independent) - using \B \ C\ \B \ D\ apply blast - apply (meson \independent C\ independent_mono inf.cobounded1) - using B \C \ S\ \D \ T\ apply auto - done - then have Deq: "D = B \ (D - C)" - by blast - have CUD: "C \ D \ {x + y |x y. x \ S \ y \ T}" - apply safe - apply (metis add.right_neutral subsetCE \C \ S\ \subspace T\ set_eq_subset span_0 span_minimal) - apply (metis add.left_neutral subsetCE \D \ T\ \subspace S\ set_eq_subset span_0 span_minimal) - done - have "a v = 0" if 0: "(\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) = 0" - and v: "v \ C \ (D-C)" for a v - proof - - have eq: "(\v\D - C. a v *\<^sub>R v) = - (\v\C. a v *\<^sub>R v)" - using that add_eq_0_iff by blast - have "(\v\D - C. a v *\<^sub>R v) \ S" - apply (subst eq) - apply (rule subspace_neg [OF \subspace S\]) - apply (rule subspace_sum [OF \subspace S\]) - by (meson subsetCE subspace_mul \C \ S\ \subspace S\) - moreover have "(\v\D - C. a v *\<^sub>R v) \ T" - apply (rule subspace_sum [OF \subspace T\]) - by (meson DiffD1 \D \ T\ \subspace T\ subset_eq subspace_def) - ultimately have "(\v \ D-C. a v *\<^sub>R v) \ span B" - using B by blast - then obtain e where e: "(\v\B. e v *\<^sub>R v) = (\v \ D-C. a v *\<^sub>R v)" - using span_finite [OF \finite B\] by blast - have "\c v. \(\v\C. c v *\<^sub>R v) = 0; v \ C\ \ c v = 0" - using independent_explicit \independent C\ by blast - define cc where "cc x = (if x \ B then a x + e x else a x)" for x - have [simp]: "C \ B = B" "D \ B = B" "C \ - B = C-D" "B \ (D - C) = {}" - using \B \ C\ \B \ D\ Beq by blast+ - have f2: "(\v\C \ D. e v *\<^sub>R v) = (\v\D - C. a v *\<^sub>R v)" - using Beq e by presburger - have f3: "(\v\C \ D. a v *\<^sub>R v) = (\v\C - D. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) + (\v\C \ D. a v *\<^sub>R v)" - using \finite C\ \finite D\ sum.union_diff2 by blast - have f4: "(\v\C \ (D - C). a v *\<^sub>R v) = (\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v)" - by (meson Diff_disjoint \finite C\ \finite D\ finite_Diff sum.union_disjoint) - have "(\v\C. cc v *\<^sub>R v) = 0" - using 0 f2 f3 f4 - apply (simp add: cc_def Beq if_smult \finite C\ sum.If_cases algebra_simps sum.distrib) - apply (simp add: add.commute add.left_commute diff_eq) - done - then have "\v. v \ C \ cc v = 0" - using independent_explicit \independent C\ by blast - then have C0: "\v. v \ C - B \ a v = 0" - by (simp add: cc_def Beq) meson - then have [simp]: "(\x\C - B. a x *\<^sub>R x) = 0" - by simp - have "(\x\C. a x *\<^sub>R x) = (\x\B. a x *\<^sub>R x)" - proof - - have "C - D = C - B" - using Beq by blast - then show ?thesis - using Beq \(\x\C - B. a x *\<^sub>R x) = 0\ f3 f4 by auto - qed - with 0 have Dcc0: "(\v\D. a v *\<^sub>R v) = 0" - apply (subst Deq) - by (simp add: \finite B\ \finite D\ sum_Un) - then have D0: "\v. v \ D \ a v = 0" - using independent_explicit \independent D\ by blast - show ?thesis - using v C0 D0 Beq by blast - qed - then have "independent (C \ (D - C))" - by (simp add: independent_explicit \finite C\ \finite D\ sum_Un del: Un_Diff_cancel) - then have indCUD: "independent (C \ D)" by simp - have "dim (S \ T) = card B" - by (rule dim_unique [OF B indB refl]) - moreover have "dim S = card C" - by (metis \C \ S\ \independent C\ \S \ span C\ basis_card_eq_dim) - moreover have "dim T = card D" - by (metis \D \ T\ \independent D\ \T \ span D\ basis_card_eq_dim) - moreover have "dim {x + y |x y. x \ S \ y \ T} = card(C \ D)" - apply (rule dim_unique [OF CUD _ indCUD refl], clarify) - apply (meson \S \ span C\ \T \ span D\ span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff) - done - ultimately show ?thesis - using \B = C \ D\ [symmetric] - by (simp add: \independent C\ \independent D\ card_Un_Int independent_finite) -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 (op+ (-a) ` S)" "affine (op+ (-a) ` T)" - using assms by (auto simp: affine_translation [symmetric]) - have zero: "0 \ (op+ (-a) ` S)" "0 \ (op+ (-a) ` T)" - using a assms by auto - have [simp]: "{x + y |x y. x \ op + (- a) ` S \ y \ op + (- a) ` T} = - op + (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" - by (force simp: algebra_simps scaleR_2) - have [simp]: "op + (- a) ` S \ op + (- a) ` T = op + (- a) ` (S \ T)" - by auto - show ?thesis - using aff_dim_sums_Int_0 [OF aff zero] - by (auto simp: aff_dim_translation_eq) -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 dim_Times: - fixes S :: "'a :: euclidean_space set" and T :: "'a set" - assumes "subspace S" "subspace T" - shows "dim(S \ T) = dim S + dim T" -proof - - have ss: "subspace ((\x. (x, 0)) ` S)" "subspace (Pair 0 ` T)" - by (rule subspace_linear_image, unfold_locales, auto simp: assms)+ - have "dim(S \ T) = dim({u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T})" - by (simp add: Times_eq_image_sum) - moreover have "dim ((\x. (x, 0::'a)) ` S) = dim S" "dim (Pair (0::'a) ` T) = dim T" - by (auto simp: additive.intro linear.intro linear_axioms.intro inj_on_def intro: dim_image_eq) - moreover have "dim ((\x. (x, 0)) ` S \ Pair 0 ` T) = 0" - by (subst dim_eq_0) (force simp: zero_prod_def) - ultimately show ?thesis - using dim_sums_Int [OF ss] by linarith -qed - -subsection\ Orthogonal bases, Gram-Schmidt process, and related theorems\ - -lemma span_delete_0 [simp]: "span(S - {0}) = span S" -proof - show "span (S - {0}) \ span S" - by (blast intro!: span_mono) -next - have "span S \ span(insert 0 (S - {0}))" - by (blast intro!: span_mono) - also have "... \ span(S - {0})" - using span_insert_0 by blast - finally show "span S \ span (S - {0})" . -qed - -lemma span_image_scale: - assumes "finite S" and nz: "\x. x \ S \ c x \ 0" - shows "span ((\x. c x *\<^sub>R x) ` S) = span S" -using assms -proof (induction S arbitrary: c) - case (empty c) show ?case by simp -next - case (insert x F c) - show ?case - proof (intro set_eqI iffI) - fix y - assume "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" - then show "y \ span (insert x F)" - using insert by (force simp: span_breakdown_eq) - next - fix y - assume "y \ span (insert x F)" - then show "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" - using insert - apply (clarsimp simp: span_breakdown_eq) - apply (rule_tac x="k / c x" in exI) - by simp - qed -qed - -lemma pairwise_orthogonal_independent: - assumes "pairwise orthogonal S" and "0 \ S" - shows "independent S" -proof - - have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" - using assms by (simp add: pairwise_def orthogonal_def) - have "False" if "a \ S" and a: "a \ span (S - {a})" for a - proof - - obtain T U where "T \ S - {a}" "a = (\v\T. U v *\<^sub>R v)" - using a by (force simp: span_explicit) - then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" - by simp - also have "... = 0" - apply (simp add: inner_sum_right) - apply (rule comm_monoid_add_class.sum.neutral) - by (metis "0" DiffE \T \ S - {a}\ mult_not_zero singletonI subsetCE \a \ S\) - finally show ?thesis - using \0 \ S\ \a \ S\ by auto - qed - then show ?thesis - by (force simp: dependent_def) -qed - -lemma pairwise_orthogonal_imp_finite: - fixes S :: "'a::euclidean_space set" - assumes "pairwise orthogonal S" - shows "finite S" -proof - - have "independent (S - {0})" - apply (rule pairwise_orthogonal_independent) - apply (metis Diff_iff assms pairwise_def) - by blast - then show ?thesis - by (meson independent_imp_finite infinite_remove) -qed - -lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" - by (simp add: subspace_def orthogonal_clauses) - -lemma subspace_orthogonal_to_vectors: "subspace {y. \x \ S. orthogonal x y}" - by (simp add: subspace_def orthogonal_clauses) - -lemma orthogonal_to_span: - assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" - shows "orthogonal x a" -apply (rule span_induct [OF a subspace_orthogonal_to_vector]) -apply (simp add: x) -done - -proposition Gram_Schmidt_step: - fixes S :: "'a::euclidean_space set" - assumes S: "pairwise orthogonal S" and x: "x \ span S" - shows "orthogonal x (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b))" -proof - - have "finite S" - by (simp add: S pairwise_orthogonal_imp_finite) - have "orthogonal (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)) x" - if "x \ S" for x - proof - - have "a \ x = (\y\S. if y = x then y \ a else 0)" - by (simp add: \finite S\ inner_commute sum.delta that) - also have "... = (\b\S. b \ a * (b \ x) / (b \ b))" - apply (rule sum.cong [OF refl], simp) - by (meson S orthogonal_def pairwise_def that) - finally show ?thesis - by (simp add: orthogonal_def algebra_simps inner_sum_left) - qed - then show ?thesis - using orthogonal_to_span orthogonal_commute x by blast -qed - - -lemma orthogonal_extension_aux: - fixes S :: "'a::euclidean_space set" - assumes "finite T" "finite S" "pairwise orthogonal S" - shows "\U. pairwise orthogonal (S \ U) \ span (S \ U) = span (S \ T)" -using assms -proof (induction arbitrary: S) - case empty then show ?case - by simp (metis sup_bot_right) -next - case (insert a T) - have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" - using insert by (simp add: pairwise_def orthogonal_def) - define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" - obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" - and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" - apply (rule exE [OF insert.IH [of "insert a' S"]]) - apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) - done - have orthS: "\x. x \ S \ a' \ x = 0" - apply (simp add: a'_def) - using Gram_Schmidt_step [OF \pairwise orthogonal S\] - apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD]) - done - have "span (S \ insert a' U) = span (insert a' (S \ T))" - using spanU by simp - also have "... = span (insert a (S \ T))" - apply (rule eq_span_insert_eq) - apply (simp add: a'_def span_neg span_sum span_clauses(1) span_mul) - done - also have "... = span (S \ insert a T)" - by simp - finally show ?case - apply (rule_tac x="insert a' U" in exI) - using orthU apply auto - done -qed - - -proposition orthogonal_extension: - fixes S :: "'a::euclidean_space set" - assumes S: "pairwise orthogonal S" - obtains U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" -proof - - obtain B where "finite B" "span B = span T" - using basis_subspace_exists [of "span T"] subspace_span by auto - with orthogonal_extension_aux [of B S] - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" - using assms pairwise_orthogonal_imp_finite by auto - show ?thesis - apply (rule_tac U=U in that) - apply (simp add: \pairwise orthogonal (S \ U)\) - by (metis \span (S \ U) = span (S \ B)\ \span B = span T\ span_Un) -qed - -corollary orthogonal_extension_strong: - fixes S :: "'a::euclidean_space set" - assumes S: "pairwise orthogonal S" - obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" - "span (S \ U) = span (S \ T)" -proof - - obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" - using orthogonal_extension assms by blast - then show ?thesis - apply (rule_tac U = "U - (insert 0 S)" in that) - apply blast - apply (force simp: pairwise_def) - apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un) - done -qed - -subsection\Decomposing a vector into parts in orthogonal subspaces.\ - -text\existence of orthonormal basis for a subspace.\ - -lemma orthogonal_spanningset_subspace: - fixes S :: "'a :: euclidean_space set" - assumes "subspace S" - obtains B where "B \ S" "pairwise orthogonal B" "span B = S" -proof - - obtain B where "B \ S" "independent B" "S \ span B" "card B = dim S" - using basis_exists by blast - with orthogonal_extension [of "{}" B] - show ?thesis - by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that) -qed - -lemma orthogonal_basis_subspace: - fixes S :: "'a :: euclidean_space set" - assumes "subspace S" - obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" - "card B = dim S" "span B = S" -proof - - obtain B where "B \ S" "pairwise orthogonal B" "span B = S" - using assms orthogonal_spanningset_subspace by blast - then show ?thesis - apply (rule_tac B = "B - {0}" in that) - apply (auto simp: indep_card_eq_dim_span pairwise_subset Diff_subset pairwise_orthogonal_independent elim: pairwise_subset) - done -qed - -proposition orthonormal_basis_subspace: - fixes S :: "'a :: euclidean_space set" - assumes "subspace S" - obtains B where "B \ S" "pairwise orthogonal B" - and "\x. x \ B \ norm x = 1" - and "independent B" "card B = dim S" "span B = S" -proof - - obtain B where "0 \ B" "B \ S" - and orth: "pairwise orthogonal B" - and "independent B" "card B = dim S" "span B = S" - by (blast intro: orthogonal_basis_subspace [OF assms]) - have 1: "(\x. x /\<^sub>R norm x) ` B \ S" - using \span B = S\ span_clauses(1) span_mul by fastforce - have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" - using orth by (force simp: pairwise_def orthogonal_clauses) - have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" - by (metis (no_types, lifting) \0 \ B\ image_iff norm_sgn sgn_div_norm) - have 4: "independent ((\x. x /\<^sub>R norm x) ` B)" - by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) - have "inj_on (\x. x /\<^sub>R norm x) B" - proof - fix x y - assume "x \ B" "y \ B" "x /\<^sub>R norm x = y /\<^sub>R norm y" - moreover have "\i. i \ B \ norm (i /\<^sub>R norm i) = 1" - using 3 by blast - ultimately show "x = y" - by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) - qed - then have 5: "card ((\x. x /\<^sub>R norm x) ` B) = dim S" - by (metis \card B = dim S\ card_image) - have 6: "span ((\x. x /\<^sub>R norm x) ` B) = S" - by (metis "1" "4" "5" assms card_eq_dim independent_finite span_subspace) - show ?thesis - by (rule that [OF 1 2 3 4 5 6]) -qed - -proposition orthogonal_subspace_decomp_exists: - fixes S :: "'a :: euclidean_space set" - obtains y z where "y \ span S" "\w. w \ span S \ orthogonal z w" "x = y + z" -proof - - obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" - using orthogonal_basis_subspace subspace_span by blast - let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" - have orth: "orthogonal (x - ?a) w" if "w \ span S" for w - by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) - show ?thesis - apply (rule_tac y = "?a" and z = "x - ?a" in that) - apply (meson \T \ span S\ span_mul span_sum subsetCE) - apply (fact orth, simp) - done -qed - -lemma orthogonal_subspace_decomp_unique: - fixes S :: "'a :: euclidean_space set" - assumes "x + y = x' + y'" - and ST: "x \ span S" "x' \ span S" "y \ span T" "y' \ span T" - and orth: "\a b. \a \ S; b \ T\ \ orthogonal a b" - shows "x = x' \ y = y'" -proof - - have "x + y - y' = x'" - by (simp add: assms) - moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" - by (meson orth orthogonal_commute orthogonal_to_span) - ultimately have "0 = x' - x" - by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) - with assms show ?thesis by auto -qed - -proposition dim_orthogonal_sum: - fixes A :: "'a::euclidean_space set" - assumes "\x y. \x \ A; y \ B\ \ x \ y = 0" - shows "dim(A \ B) = dim A + dim B" -proof - - have 1: "\x y. \x \ span A; y \ B\ \ x \ y = 0" - by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) - have "\x y. \x \ span A; y \ span B\ \ x \ y = 0" - apply (erule span_induct [OF _ subspace_hyperplane]) - using 1 by (simp add: ) - then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" - by simp - have "dim(A \ B) = dim (span (A \ B))" - by simp - also have "... = dim ((\(a, b). a + b) ` (span A \ span B))" - by (simp add: span_Un) - also have "... = dim {x + y |x y. x \ span A \ y \ span B}" - by (auto intro!: arg_cong [where f=dim]) - also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" - by (auto simp: dest: 0) - also have "... = dim (span A) + dim (span B)" - by (rule dim_sums_Int) auto - also have "... = dim A + dim B" - by simp - finally show ?thesis . -qed - -lemma dim_subspace_orthogonal_to_vectors: - fixes A :: "'a::euclidean_space set" - assumes "subspace A" "subspace B" "A \ B" - shows "dim {y \ B. \x \ A. orthogonal x y} + dim A = dim B" -proof - - have "dim (span ({y \ B. \x\A. orthogonal x y} \ A)) = dim (span B)" - proof (rule arg_cong [where f=dim, OF subset_antisym]) - show "span ({y \ B. \x\A. orthogonal x y} \ A) \ span B" - by (simp add: \A \ B\ Collect_restrict span_mono) - next - have *: "x \ span ({y \ B. \x\A. orthogonal x y} \ A)" - if "x \ B" for x - proof - - obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" - using orthogonal_subspace_decomp_exists [of A x] that by auto - have "y \ span B" - by (metis span_eq \y \ span A\ assms subset_iff) - then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" - by simp (metis (no_types) span_eq \x = y + z\ \subspace A\ \subspace B\ orth orthogonal_commute span_add_eq that) - then have z: "z \ span {y \ B. \x\A. orthogonal x y}" - by (meson span_inc subset_iff) - then show ?thesis - apply (simp add: span_Un image_def) - apply (rule bexI [OF _ z]) - apply (simp add: \x = y + z\ \y \ span A\) - done - qed - show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" - by (rule span_minimal) (auto intro: * span_minimal elim: ) - qed - then show ?thesis - by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) -qed - -lemma aff_dim_openin: - fixes S :: "'a::euclidean_space set" - assumes ope: "openin (subtopology euclidean 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 \independent B\ - apply (rule independent_injective_image, simp) - by (metis \0 < e\ injective_scaleR less_irrefl) - qed - also have "... = aff_dim S" - using \a \ S\ aff_dim_eq_dim hull_inc by force - finally show "aff_dim T \ aff_dim S" . - qed -qed - -lemma dim_openin: - fixes S :: "'a::euclidean_space set" - assumes ope: "openin (subtopology euclidean 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_inc subspace_span) - finally show "dim T \ dim S" by simp -qed - -subsection\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 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 (op + (- z) ` S)" by blast - then have aff: "affine (op + (- z) ` S)" and span: "span (op + (- z) ` S) = (op + (- z) ` S)" - by (auto simp: subspace_imp_affine) - obtain a' a'' where a': "a' \ span (op + (- z) ` S)" and a: "a = a' + a''" - and "\w. w \ span (op + (- z) ` S) \ orthogonal a'' w" - using orthogonal_subspace_decomp_exists [of "op + (- z) ` S" "a"] by metis - then have "\w. w \ S \ a'' \ (w-z) = 0" - by (simp add: imageI orthogonal_def span) - 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 \ op + (- z) ` S \ (w + a') \ op + (- z) ` S" - by (metis subspace_add a' span_eq 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\) - 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) - 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 linear_injective_left_inverse assms 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 - - -text\Choosing a subspace of a given dimension\ -proposition choose_subspace_of_subspace: - fixes S :: "'n::euclidean_space set" - assumes "n \ dim S" - obtains T where "subspace T" "T \ span S" "dim T = n" -proof - - have "\T. subspace T \ T \ span S \ dim T = n" - using assms - proof (induction n) - case 0 then show ?case by force - next - case (Suc n) - then obtain T where "subspace T" "T \ span S" "dim T = n" - by force - then show ?case - proof (cases "span S \ span T") - case True - have "dim S = dim T" - apply (rule span_eq_dim [OF subset_antisym [OF True]]) - by (simp add: \T \ span S\ span_minimal subspace_span) - then show ?thesis - using Suc.prems \dim T = n\ by linarith - next - case False - then obtain y where y: "y \ S" "y \ T" - by (meson span_mono subsetI) - then have "span (insert y T) \ span S" - by (metis (no_types) \T \ span S\ subsetD insert_subset span_inc span_mono span_span) - with \dim T = n\ \subspace T\ y show ?thesis - apply (rule_tac x="span(insert y T)" in exI) - apply (auto simp: dim_insert) - using span_eq by blast - qed - qed - with that show ?thesis by blast -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 (op + (- a) ` S)" - by (simp add: affine_diffs_subspace) - have "nat d \ dim (op + (- 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 (op + (- a) ` S)" - and Tdim: "dim T = nat d" - using choose_subspace_of_subspace [of "nat d" "op + (- a) ` S"] by blast - then have "affine T" - using subspace_affine by blast - then have "affine (op + a ` T)" - by (metis affine_hull_eq affine_hull_translation) - moreover have "op + a ` T \ S" - proof - - have "T \ op + (- a) ` S" - by (metis (no_types) span_eq Tsb ss) - then show "op + a ` T \ S" - using add_ac by auto - qed - moreover have "aff_dim (op + 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\Several Variants of Paracompactness\ - -proposition paracompact: - fixes S :: "'a :: euclidean_space 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) - done - 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 \" "\\ = UNION S G" - using Lindelof [of "G ` S"] by (metis image_iff) - then obtain K where K: "K \ S" "countable K" and eq: "UNION K G = UNION S G" - 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 :: euclidean_space set" - assumes cin: "closedin (subtopology euclidean U) S" - and oin: "\T. T \ \ \ openin (subtopology euclidean U) T" - and "S \ \\" - obtains \' where "S \ \ \'" - and "\V. V \ \' \ openin (subtopology euclidean U) V \ (\T. T \ \ \ V \ T)" - and "\x. x \ U - \ \V. openin (subtopology euclidean 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 \ {}}" - using paracompact [OF 1 2] by auto - 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 (subtopology euclidean 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 (subtopology euclidean 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 paracompact_closed: - fixes S :: "'a :: euclidean_space 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 \ {})}" -using paracompact_closedin [of UNIV S \] assms -by (auto simp: open_openin [symmetric] closed_closedin [symmetric]) - - -subsection\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 (subtopology euclidean (S \ T)) ((\x. Pair x (f x)) ` S)" -proof - - have eq: "((\x. Pair x (f x)) ` S) = {z. z \ S \ T \ (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 (subtopology euclidean (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 (subtopology euclidean T) U" - have eq: "{x. x \ S \ f x \ U} = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" - by (force simp: image_iff) - show "closedin (subtopology euclidean S) {x \ S. f x \ 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 (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (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}" "S \ T" "\i. i" "\i. f" f] contf contg opS opT by blast - -lemma continuous_on_cases_local_open: - assumes opS: "openin (subtopology euclidean (S \ T)) S" - and opT: "openin (subtopology euclidean (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 - -subsection\The union of two collinear segments is another segment\ - -proposition 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 real_vector.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 real_vector.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\ \UNION UNIV f = 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 divide_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 \ UNION UNIV ?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: divide_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 \ UNION UNIV ?C" - by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) - qed - then show "UNION UNIV ?C = S" by auto - qed - ultimately show ?thesis - using that by metis -qed - end diff -r e5995950b98a -r 2562f151541c src/HOL/Analysis/Starlike.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Analysis/Starlike.thy Thu Jul 20 14:05:29 2017 +0100 @@ -0,0 +1,7896 @@ +(* 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 +*) + +section \Line segments, Starlike Sets, etc.\ + +theory Starlike + imports Convex_Euclidean_Space + +begin + +definition midpoint :: "'a::real_vector \ 'a \ 'a" + where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" + +definition 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 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: + "linear f \ closed_segment (f a) (f b) = f ` (closed_segment a b)" + by (force simp add: in_segment linear_add_cmul) + +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 midpoint_idem [simp]: "midpoint x x = x" + unfolding midpoint_def + unfolding scaleR_right_distrib + unfolding scaleR_left_distrib[symmetric] + by auto + +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\Starlike sets\ + +definition "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" + +lemma starlike_UNIV [simp]: "starlike UNIV" + by (simp add: starlike_def) + +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_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 convex_imp_starlike: + "convex S \ S \ {} \ starlike S" + unfolding convex_contains_segment starlike_def by auto + +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 real_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]: + fixes a :: "'a::euclidean_space" + shows "closure(open_segment a b) = (if a = b then {} else closed_segment a b)" +proof - + have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" if "a \ b" + apply (rule closure_injective_linear_image [symmetric]) + apply (simp add:) + using that by (simp add: inj_on_def) + then show ?thesis + by (simp add: segment_image_interval segment_eq_compose closure_greaterThanLessThan [symmetric] + closure_translation image_comp [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 show ?thesis + apply (simp add: open_segment_image_interval segment_eq_compose) + by (metis image_comp convex_translation) +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 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\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 "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 + unfolding between_def split_conv + by (auto simp add: 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) + show ?thesis + unfolding between_def split_conv closed_segment_def mem_Collect_eq + apply rule + apply (elim exE conjE) + apply (subst dist_triangle_eq) + proof - + fix u + assume as: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" + then have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" + unfolding as(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 as(2,3) + by (auto simp add: field_simps) + next + assume as: "dist a b = dist a x + dist x b" + have "norm (a - x) / norm (a - b) \ 1" + using Fal2 unfolding as[unfolded dist_norm] norm_ge_zero by auto + then show "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" + apply (rule_tac x="dist a x / dist a b" in exI) + unfolding dist_norm + apply (subst euclidean_eq_iff) + apply rule + defer + apply rule + prefer 3 + apply rule + proof - + fix i :: 'a + assume i: "i \ Basis" + have "((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^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 as[unfolded dist_norm] + using as[unfolded dist_triangle_eq] + apply - + apply (subst (asm) euclidean_eq_iff) + using i + apply (erule_tac x=i in ballE) + apply (auto simp add: field_simps inner_simps) + done + finally show "x \ i = + ((1 - norm (a - x) / norm (a - b)) *\<^sub>R a + (norm (a - x) / norm (a - b)) *\<^sub>R b) \ i" + by auto + qed (insert Fal2, auto) + qed +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 + apply(rule_tac[!] *) + unfolding euclidean_eq_iff[where 'a='a] + apply (auto simp add: field_simps inner_simps) + done +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 \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 + apply (rule_tac x="e*d" in exI) + apply rule + defer + unfolding subset_eq Ball_def mem_ball + proof (rule, rule) + fix y + assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" + 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) + apply (auto simp add:field_simps norm_minus_commute) + done + then show ?thesis + unfolding * + apply - + apply (rule mem_interior_convex_shrink) + using assms(1,4-5) \y\s\ + apply auto + done +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: divide_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 \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)}" + unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]] + apply (rule set_eqI, rule) + unfolding mem_Collect_eq + apply (erule_tac[!] exE) + apply (erule_tac[!] conjE)+ + unfolding sum_clauses(2)[OF \finite s\] + apply (rule_tac x=u in exI) + defer + apply (rule_tac x="\x. if x = 0 then 1 - sum u s else u x" in exI) + using assms(2) + unfolding if_smult and sum_delta_notmem[OF assms(2)] + apply auto + done + +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\] + apply (rule set_eqI) + unfolding mem_Collect_eq + apply rule + apply (elim exE conjE) + apply (erule_tac[2] conjE)+ + proof - + fix x :: "'a::euclidean_space" + fix u + assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" "(\x\?D. u x *\<^sub>R x) = x" + have *: "\i\Basis. i:d \ u i = x\i" + and "(\i\Basis. i \ d \ x \ i = 0)" + using as(3) + unfolding substdbasis_expansion_unique[OF assms] + by auto + then have **: "sum u ?D = sum (op \ x) ?D" + apply - + apply (rule sum.cong) + using assms + apply auto + done + have "(\i\Basis. 0 \ x\i) \ sum (op \ x) ?D \ 1" + proof (rule,rule) + fix i :: 'a + assume i: "i \ Basis" + have "i \ d \ 0 \ x\i" + unfolding *[rule_format,OF i,symmetric] + apply (rule_tac as(1)[rule_format]) + apply auto + done + moreover have "i \ d \ 0 \ x\i" + using \(\i\Basis. i \ d \ x \ i = 0)\[rule_format, OF i] by auto + ultimately show "0 \ x\i" by auto + qed (insert as(2)[unfolded **], auto) + then show "(\i\Basis. 0 \ x\i) \ sum (op \ x) ?D \ 1 \ (\i\Basis. i \ d \ x \ i = 0)" + using \(\i\Basis. i \ d \ x \ i = 0)\ by auto + next + fix x :: "'a::euclidean_space" + assume as: "\i\Basis. 0 \ x \ i" "sum (op \ x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" + show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" + using as d + unfolding substdbasis_expansion_unique[OF assms] + apply (rule_tac x="inner x" in exI) + apply auto + done + 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}" + apply (rule set_eqI) + unfolding mem_interior std_simplex + unfolding subset_eq mem_Collect_eq Ball_def mem_ball + unfolding Ball_def[symmetric] + apply rule + apply (elim exE conjE) + defer + apply (erule conjE) +proof - + fix x :: 'a + fix e + assume "e > 0" and as: "\xa. dist x xa < e \ (\x\Basis. 0 \ xa \ x) \ sum (op \ xa) Basis \ 1" + show "(\xa\Basis. 0 < x \ xa) \ sum (op \ x) Basis < 1" + apply safe + proof - + fix i :: 'a + assume i: "i \ Basis" + then show "0 < x \ i" + using as[THEN spec[where x="x - (e / 2) *\<^sub>R i"]] and \e > 0\ + unfolding dist_norm + by (auto elim!: ballE[where x=i] simp: 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 (op \ (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" + apply (rule_tac sum.cong) + apply auto + done + have "sum (op \ x) Basis < sum (op \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" + unfolding * sum.distrib + using \e > 0\ DIM_positive[where 'a='a] + apply (subst sum.delta') + apply (auto simp: SOME_Basis) + done + also have "\ \ 1" + using ** + apply (drule_tac as[rule_format]) + apply auto + done + finally show "sum (op \ x) Basis < 1" by auto + qed +next + fix x :: 'a + assume as: "\i\Basis. 0 < x \ i" "sum (op \ x) Basis < 1" + obtain a :: 'b where "a \ UNIV" using UNIV_witness .. + let ?d = "(1 - sum (op \ x) Basis) / real (DIM('a))" + show "\e>0. \y. dist x y < e \ (\i\Basis. 0 \ y \ i) \ sum (op \ y) Basis \ 1" + proof (rule_tac x="min (Min ((op \ x) ` Basis)) D" for D in exI, intro conjI impI allI) + fix y + assume y: "dist x y < min (Min (op \ x ` Basis)) ?d" + have "sum (op \ y) Basis \ sum (\i. x\i + ?d) Basis" + proof (rule sum_mono) + fix i :: 'a + assume i: "i \ Basis" + then have "\y\i - x\i\ < ?d" + apply - + apply (rule le_less_trans) + using Basis_le_norm[OF i, of "y - x"] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] + apply (auto simp add: norm_minus_commute inner_diff_left) + done + 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 (op \ y) Basis \ 1" . + show "(\i\Basis. 0 \ y \ i)" + proof safe + fix i :: 'a + assume i: "i \ Basis" + have "norm (x - y) < x\i" + apply (rule less_le_trans) + apply (rule y[unfolded min_less_iff_conj dist_norm, THEN conjunct1]) + using i + apply auto + done + 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 ((op \ x) ` Basis) > 0" + using as by simp + moreover have "?d > 0" + using as by (auto simp: Suc_le_eq) + ultimately show "0 < min (Min (op \ x ` Basis)) ((1 - sum (op \ 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 (op \ ?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 (op \ ?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" + apply (rule finite_subset) + using assms + apply auto + done + 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 e0: "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: "\xa. dist x xa < e \ (\i\Basis. i \ d \ xa\i = 0) \ + (\i\d. 0 \ xa \ i) \ sum (op \ xa) 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 (op \ x) d < 1 \ (\i\Basis. i \ d \ x\i = 0)" + apply rule + apply rule + proof - + fix i :: 'a + assume "i \ d" + then have "\ia\d. 0 \ (x - (e / 2) *\<^sub>R i) \ ia" + apply - + apply (rule as[rule_format,THEN conjunct1]) + unfolding dist_norm + using d \e > 0\ x0 + apply (auto simp: inner_simps inner_Basis) + done + then show "0 < x \ i" + apply (erule_tac x=i in ballE) + using \e > 0\ \i \ d\ d + apply (auto simp: inner_simps inner_Basis) + done + 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 (op \ (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 (op \ x) d < sum (op \ (x + (e / 2) *\<^sub>R a)) d" + unfolding * sum.distrib + using \e > 0\ \a \ d\ + using \finite d\ + by (auto simp add: sum.delta') + also have "\ \ 1" + using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] + by auto + finally show "sum (op \ x) d < 1 \ (\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 (op \ x) d) / real (card d)" + have "0 < card d" using \d \ {}\ \finite d\ + by (simp add: card_gt_0_iff) + have "Min ((op \ 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 ((op \ 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 ((op \ 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 (op \ x ` d)) ?d" + assume y2: "\i\Basis. i \ d \ y\i = 0" + have "sum (op \ 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\ < ?d" + apply (rule le_less_trans) + using Basis_le_norm[OF i, of "y - x"] + using y[unfolded min_less_iff_conj dist_norm, THEN conjunct2] + apply (auto simp add: norm_minus_commute inner_simps) + done + 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 (op \ 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 "op \ 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 (op \ 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 set_rev_mp[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 (op \ ?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 (op \ ?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_superset[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 \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 auto + 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_inc 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 \ op + (-a) ` S" + using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto + then have "rel_interior (op + (-a) ` S) \ {}" + using rel_interior_convex_nonempty_aux[of "op + (-a) ` S"] + convex_translation[of S "-a"] assms + by auto + then have "rel_interior S \ {}" + using rel_interior_translation by auto + } + 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" + apply (simp add: hull_inc affine_hull_span_0) + using assms dim_eq_full indep_card_eq_dim_span by fastforce + 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 (subtopology euclidean (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)" + apply auto + using assms + apply simp + done + 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 "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" + apply (simp add: rel_frontier_def) + apply (simp add: rel_interior_eq_closure [symmetric]) + using rel_interior_subset_closure by blast + +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 (subtopology euclidean (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 \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 + 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 \linear f\ by (simp add: linear_add_cmul) + 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 - + 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 (metis assms(1) assms(2) convex_Int subspace_UNIV subspace_imp_convex subspace_linear_image) + { + 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 (metis assms(1) subspace_UNIV subspace_imp_affine subspace_linear_image) + 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 "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^sub>R c) ` S)" + using rel_interior_injective_linear_image[of "(op *\<^sub>R c)" S] + linear_conv_bounded_linear[of "op *\<^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 "(op *\<^sub>R c) ` (rel_interior S) = rel_interior ((op *\<^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 ((op *\<^sub>R c) ` S) \ rel_open ((op *\<^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 \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 \ op *\<^sub>R c ` S = S)" + using cone_iff[of S] assms by auto + then have *: "0 \ ({0} \ rel_interior S)" + and "\c. c > 0 \ op *\<^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 \ ((op *\<^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 = (op *\<^sub>R c ` S)" + using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto + then have "rel_interior (f c) = op *\<^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 + have "z \ ?rhs" + using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ + apply auto + apply (rule_tac x = "fst z" in exI) + apply (rule_tac x = x in exI) + using * + apply auto + done + } + 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 + + +subsection \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 + apply auto + apply (rule_tac x = c in exI) + apply (rule_tac x = s in exI) + using cs + apply auto + done + } + 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 simp add: convex_convex_hull) + } + 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\Explicit formulas for interior and relative interior of convex hull\ + +lemma interior_atLeastAtMost [simp]: + fixes a::real shows "interior {a..b} = {a<.. x < b \ (at x within {a..b}) = at x" + by (metis at_within_interior greaterThanLessThan_iff interior_atLeastAtMost) + +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) + apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) + apply (metis Suc_le_lessD not_less order_trans card_image_le finite_imageI dim_le_card) + done + +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_hull_convex_hull 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\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 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\The infimum of the distance between two sets\ + +definition setdist :: "'a::metric_space set \ 'a set \ real" where + "setdist s t \ + (if s = {} \ t = {} then 0 + else Inf {dist x y| x y. x \ s \ y \ t})" + +lemma setdist_empty1 [simp]: "setdist {} t = 0" + by (simp add: setdist_def) + +lemma setdist_empty2 [simp]: "setdist t {} = 0" + by (simp add: setdist_def) + +lemma setdist_pos_le [simp]: "0 \ setdist s t" + by (auto simp: setdist_def ex_in_conv [symmetric] intro: cInf_greatest) + +lemma le_setdistI: + assumes "s \ {}" "t \ {}" "\x y. \x \ s; y \ t\ \ d \ dist x y" + shows "d \ setdist s t" + using assms + by (auto simp: setdist_def Set.ex_in_conv [symmetric] intro: cInf_greatest) + +lemma setdist_le_dist: "\x \ s; y \ t\ \ setdist s t \ dist x y" + unfolding setdist_def + by (auto intro!: bdd_belowI [where m=0] cInf_lower) + +lemma le_setdist_iff: + "d \ setdist s t \ + (\x \ s. \y \ t. d \ dist x y) \ (s = {} \ t = {} \ d \ 0)" + apply (cases "s = {} \ t = {}") + apply (force simp add: setdist_def) + apply (intro iffI conjI) + using setdist_le_dist apply fastforce + apply (auto simp: intro: le_setdistI) + done + +lemma setdist_ltE: + assumes "setdist s t < b" "s \ {}" "t \ {}" + obtains x y where "x \ s" "y \ t" "dist x y < b" +using assms +by (auto simp: not_le [symmetric] le_setdist_iff) + +lemma setdist_refl: "setdist s s = 0" + apply (cases "s = {}") + apply (force simp add: setdist_def) + apply (rule antisym [OF _ setdist_pos_le]) + apply (metis all_not_in_conv dist_self setdist_le_dist) + done + +lemma setdist_sym: "setdist s t = setdist t s" + by (force simp: setdist_def dist_commute intro!: arg_cong [where f=Inf]) + +lemma setdist_triangle: "setdist s t \ setdist s {a} + setdist {a} t" +proof (cases "s = {} \ t = {}") + case True then show ?thesis + using setdist_pos_le by fastforce +next + case False + have "\x. x \ s \ setdist s t - dist x a \ setdist {a} t" + apply (rule le_setdistI, blast) + using False apply (fastforce intro: le_setdistI) + apply (simp add: algebra_simps) + apply (metis dist_commute dist_triangle3 order_trans [OF setdist_le_dist]) + done + then have "setdist s t - setdist {a} t \ setdist s {a}" + using False by (fastforce intro: le_setdistI) + then show ?thesis + by (simp add: algebra_simps) +qed + +lemma setdist_singletons [simp]: "setdist {x} {y} = dist x y" + by (simp add: setdist_def) + +lemma setdist_Lipschitz: "\setdist {x} s - setdist {y} s\ \ dist x y" + apply (subst setdist_singletons [symmetric]) + by (metis abs_diff_le_iff diff_le_eq setdist_triangle setdist_sym) + +lemma continuous_at_setdist [continuous_intros]: "continuous (at x) (\y. (setdist {y} s))" + by (force simp: continuous_at_eps_delta dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) + +lemma continuous_on_setdist [continuous_intros]: "continuous_on t (\y. (setdist {y} s))" + by (metis continuous_at_setdist continuous_at_imp_continuous_on) + +lemma uniformly_continuous_on_setdist: "uniformly_continuous_on t (\y. (setdist {y} s))" + by (force simp: uniformly_continuous_on_def dist_real_def intro: le_less_trans [OF setdist_Lipschitz]) + +lemma setdist_subset_right: "\t \ {}; t \ u\ \ setdist s u \ setdist s t" + apply (cases "s = {} \ u = {}", force) + apply (auto simp: setdist_def intro!: bdd_belowI [where m=0] cInf_superset_mono) + done + +lemma setdist_subset_left: "\s \ {}; s \ t\ \ setdist t u \ setdist s u" + by (metis setdist_subset_right setdist_sym) + +lemma setdist_closure_1 [simp]: "setdist (closure s) t = setdist s t" +proof (cases "s = {} \ t = {}") + case True then show ?thesis by force +next + case False + { fix y + assume "y \ t" + have "continuous_on (closure s) (\a. dist a y)" + by (auto simp: continuous_intros dist_norm) + then have *: "\x. x \ closure s \ setdist s t \ dist x y" + apply (rule continuous_ge_on_closure) + apply assumption + apply (blast intro: setdist_le_dist \y \ t\ ) + done + } note * = this + show ?thesis + apply (rule antisym) + using False closure_subset apply (blast intro: setdist_subset_left) + using False * + apply (force simp add: closure_eq_empty intro!: le_setdistI) + done +qed + +lemma setdist_closure_2 [simp]: "setdist t (closure s) = setdist t s" +by (metis setdist_closure_1 setdist_sym) + +lemma setdist_compact_closed: + fixes S :: "'a::euclidean_space set" + assumes S: "compact S" and T: "closed T" + and "S \ {}" "T \ {}" + shows "\x \ S. \y \ T. dist x y = setdist S T" +proof - + have "(\x\ S. \y \ T. {x - y}) \ {}" + using assms by blast + then have "\x \ S. \y \ T. dist x y \ setdist S T" + apply (rule distance_attains_inf [where a=0, OF compact_closed_differences [OF S T]]) + apply (simp add: dist_norm le_setdist_iff) + apply blast + done + then show ?thesis + by (blast intro!: antisym [OF _ setdist_le_dist] ) +qed + +lemma setdist_closed_compact: + fixes S :: "'a::euclidean_space set" + assumes S: "closed S" and T: "compact T" + and "S \ {}" "T \ {}" + shows "\x \ S. \y \ T. dist x y = setdist S T" + using setdist_compact_closed [OF T S \T \ {}\ \S \ {}\] + by (metis dist_commute setdist_sym) + +lemma setdist_eq_0I: "\x \ S; x \ T\ \ setdist S T = 0" + by (metis antisym dist_self setdist_le_dist setdist_pos_le) + +lemma setdist_eq_0_compact_closed: + fixes S :: "'a::euclidean_space set" + assumes S: "compact S" and T: "closed T" + shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" + apply (cases "S = {} \ T = {}", force) + using setdist_compact_closed [OF S T] + apply (force intro: setdist_eq_0I ) + done + +corollary setdist_gt_0_compact_closed: + fixes S :: "'a::euclidean_space set" + assumes S: "compact S" and T: "closed T" + shows "setdist S T > 0 \ (S \ {} \ T \ {} \ S \ T = {})" + using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] + by linarith + +lemma setdist_eq_0_closed_compact: + fixes S :: "'a::euclidean_space set" + assumes S: "closed S" and T: "compact T" + shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" + using setdist_eq_0_compact_closed [OF T S] + by (metis Int_commute setdist_sym) + +lemma setdist_eq_0_bounded: + fixes S :: "'a::euclidean_space set" + assumes "bounded S \ bounded T" + shows "setdist S T = 0 \ S = {} \ T = {} \ closure S \ closure T \ {}" + apply (cases "S = {} \ T = {}", force) + using setdist_eq_0_compact_closed [of "closure S" "closure T"] + setdist_eq_0_closed_compact [of "closure S" "closure T"] assms + apply (force simp add: bounded_closure compact_eq_bounded_closed) + done + +lemma setdist_unique: + "\a \ S; b \ T; \x y. x \ S \ y \ T ==> dist a b \ dist x y\ + \ setdist S T = dist a b" + by (force simp add: setdist_le_dist le_setdist_iff intro: antisym) + +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 setdist_eq_0_sing_1: + fixes S :: "'a::euclidean_space set" + shows "setdist {x} S = 0 \ S = {} \ x \ closure S" + by (auto simp: setdist_eq_0_bounded) + +lemma setdist_eq_0_sing_2: + fixes S :: "'a::euclidean_space set" + shows "setdist S {x} = 0 \ S = {} \ x \ closure S" + by (auto simp: setdist_eq_0_bounded) + +lemma setdist_neq_0_sing_1: + fixes S :: "'a::euclidean_space set" + shows "\setdist {x} S = a; a \ 0\ \ S \ {} \ x \ closure S" + by (auto simp: setdist_eq_0_sing_1) + +lemma setdist_neq_0_sing_2: + fixes S :: "'a::euclidean_space set" + shows "\setdist S {x} = a; a \ 0\ \ S \ {} \ x \ closure S" + by (auto simp: setdist_eq_0_sing_2) + +lemma setdist_sing_in_set: + fixes S :: "'a::euclidean_space set" + shows "x \ S \ setdist {x} S = 0" + using closure_subset by (auto simp: setdist_eq_0_sing_1) + +lemma setdist_le_sing: "x \ S ==> setdist S T \ setdist {x} T" + using setdist_subset_left by auto + +lemma setdist_eq_0_closed: + fixes S :: "'a::euclidean_space set" + shows "closed S \ (setdist {x} S = 0 \ S = {} \ x \ S)" +by (simp add: setdist_eq_0_sing_1) + +lemma setdist_eq_0_closedin: + fixes S :: "'a::euclidean_space set" + shows "\closedin (subtopology euclidean u) S; x \ u\ + \ (setdist {x} S = 0 \ S = {} \ x \ S)" + by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) + +subsection\Basic lemmas about hyperplanes and halfspaces\ + +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\Use set distance for an easy proof of separation properties\ + +proposition 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 (subtopology euclidean U) S" + and UT: "closedin (subtopology euclidean U) T" + and "S \ T = {}" + obtains S' T' where "openin (subtopology euclidean U) S'" + "openin (subtopology euclidean U) T'" + "S \ S'" "T \ T'" "S' \ T' = {}" +proof (cases "S = {} \ T = {}") + case True with that show ?thesis + apply safe + using UT closedin_subset apply blast + using US closedin_subset apply blast + done +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' = "{x\U. f x > 0}" and T' = "{x\U. f x < 0}" in that) + show "{x \ U. 0 < f x} \ {x \ U. f x < 0} = {}" + by auto + have "openin (subtopology euclidean U) {x \ U. f x \ {0<..}}" + by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) + then show "openin (subtopology euclidean U) {x \ U. 0 < f x}" by simp + next + have "openin (subtopology euclidean U) {x \ U. f x \ {..<0}}" + by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) + then show "openin (subtopology euclidean U) {x \ U. f x < 0}" by simp + next + show "S \ {x \ U. 0 < f x}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False Int_insert_right closedin_imp_subset empty_not_insert insert_absorb insert_subset linorder_neqE_linordered_idom not_le setdist_eq_0_closedin setdist_pos_le) + show "T \ {x \ U. f x < 0}" + apply (clarsimp simp add: f_def setdist_sing_in_set) + using assms + by (metis False closedin_subset disjoint_iff_not_equal insert_absorb insert_subset linorder_neqE_linordered_idom not_less setdist_eq_0_closedin setdist_pos_le topspace_euclidean_subtopology) + 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\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 (subtopology euclidean S) K" + and com: "\U. \U \ T; compact U\ \ compact {x \ S. f x \ U}" + and "f ` S \ T" + shows "closedin (subtopology euclidean 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 \ {a \ S. f a \ 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 *: "{x \ S. f x \ 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 {x \ f ` S. g x \ C}" if C: "C \ S" "compact C" for C + proof - + obtain h :: "'a set \ 'a" where "h C \ C \ h C \ S \ compact (f ` C)" + by (force simp: C RHS) + moreover have "f ` C = {b \ f ` S. g b \ C}" + using C gf by auto + ultimately show "compact {b \ f ` S. g b \ C}" + using C by auto + qed + show ?lhs + using proper_map [OF _ _ gfim] ** + by (simp add: continuous_on_closed * closedin_imp_subset) +qed + +subsection\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 (subtopology euclidean T) K" + shows "compact {x. x \ S \ f x \ 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 {z \ S \ T. fst z \ K}" +proof - + have "{z \ S \ T. fst z \ 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 (subtopology euclidean (S \ T)) c" + shows "closedin (subtopology euclidean 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 {z \ S \ T. snd z \ K}" +proof - + have "{z \ S \ T. snd z \ 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 (subtopology euclidean (S \ T)) c" + shows "closedin (subtopology euclidean 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 (subtopology euclidean (S \ T)) U" + shows "closedin (subtopology euclidean 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 (subtopology euclidean 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\Representing affine hull as a finite intersection of hyperplanes\ + +proposition 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 \ op + 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 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_superset) + 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 (subtopology euclidean (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 (subtopology euclidean (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 (subtopology euclidean (affine hull S)) S" and "finite F" "F \ S" + shows "affine hull (S - F) = affine hull S" +proof - + have clo: "closedin (subtopology euclidean 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 + +subsection\Properties of special hyperplanes\ + +lemma subspace_hyperplane: "subspace {x. a \ x = 0}" + by (simp add: subspace_def inner_right_distrib) + +lemma subspace_hyperplane2: "subspace {x. x \ a = 0}" + by (simp add: inner_commute inner_right_distrib subspace_def) + +lemma special_hyperplane_span: + fixes S :: "'n::euclidean_space set" + assumes "k \ Basis" + shows "{x. k \ x = 0} = span (Basis - {k})" +proof - + have *: "x \ span (Basis - {k})" if "k \ x = 0" for x + proof - + have "x = (\b\Basis. (x \ b) *\<^sub>R b)" + by (simp add: euclidean_representation) + also have "... = (\b \ Basis - {k}. (x \ b) *\<^sub>R b)" + by (auto simp: sum.remove [of _ k] inner_commute assms that) + finally have "x = (\b\Basis - {k}. (x \ b) *\<^sub>R b)" . + then show ?thesis + by (simp add: Linear_Algebra.span_finite) metis + qed + show ?thesis + apply (rule span_subspace [symmetric]) + using assms + apply (auto simp: inner_not_same_Basis intro: * subspace_hyperplane) + done +qed + +lemma dim_special_hyperplane: + fixes k :: "'n::euclidean_space" + shows "k \ Basis \ dim {x. k \ x = 0} = DIM('n) - 1" +apply (simp add: special_hyperplane_span) +apply (rule Linear_Algebra.dim_unique [OF subset_refl]) +apply (auto simp: Diff_subset independent_substdbasis) +apply (metis member_remove remove_def span_clauses(1)) +done + +proposition dim_hyperplane: + fixes a :: "'a::euclidean_space" + assumes "a \ 0" + shows "dim {x. a \ x = 0} = DIM('a) - 1" +proof - + have span0: "span {x. a \ x = 0} = {x. a \ x = 0}" + by (rule span_unique) (auto simp: subspace_hyperplane) + then obtain B where "independent B" + and Bsub: "B \ {x. a \ x = 0}" + and subspB: "{x. a \ x = 0} \ span B" + and card0: "(card B = dim {x. a \ x = 0})" + and ortho: "pairwise orthogonal B" + using orthogonal_basis_exists by metis + with assms have "a \ span B" + by (metis (mono_tags, lifting) span_eq inner_eq_zero_iff mem_Collect_eq span0 span_subspace) + then have ind: "independent (insert a B)" + by (simp add: \independent B\ independent_insert) + have "finite B" + using \independent B\ independent_bound by blast + have "UNIV \ span (insert a B)" + proof fix y::'a + obtain r z where z: "y = r *\<^sub>R a + z" "a \ z = 0" + apply (rule_tac r="(a \ y) / (a \ a)" and z = "y - ((a \ y) / (a \ a)) *\<^sub>R a" in that) + using assms + by (auto simp: algebra_simps) + show "y \ span (insert a B)" + by (metis (mono_tags, lifting) z Bsub Convex_Euclidean_Space.span_eq + add_diff_cancel_left' mem_Collect_eq span0 span_breakdown_eq span_subspace subspB) + qed + then have dima: "DIM('a) = dim(insert a B)" + by (metis antisym dim_UNIV dim_subset_UNIV subset_le_dim) + then show ?thesis + by (metis (mono_tags, lifting) Bsub Diff_insert_absorb \a \ span B\ ind card0 card_Diff_singleton dim_span indep_card_eq_dim_span insertI1 subsetCE subspB) +qed + +lemma lowdim_eq_hyperplane: + fixes S :: "'a::euclidean_space set" + assumes "dim S = DIM('a) - 1" + obtains a where "a \ 0" and "span S = {x. a \ x = 0}" +proof - + have [simp]: "dim S < DIM('a)" + by (simp add: DIM_positive assms) + then obtain b where b: "b \ 0" "span S \ {a. b \ a = 0}" + using lowdim_subset_hyperplane [of S] by fastforce + show ?thesis + using b that real_vector_class.subspace_span [of S] + by (simp add: assms dim_hyperplane subspace_dim_equal subspace_hyperplane) +qed + +lemma dim_eq_hyperplane: + fixes S :: "'n::euclidean_space set" + shows "dim S = DIM('n) - 1 \ (\a. a \ 0 \ span S = {x. a \ x = 0})" +by (metis One_nat_def dim_hyperplane dim_span lowdim_eq_hyperplane) + +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 + 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 (rule ex_cong) + apply (metis (mono_tags) span_0 \c = 0\ image_add_0 inner_zero_right mem_Collect_eq) + done + next + case False + have xc_im: "x \ op + 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 \ op + c ` {y. a \ y = 0}" + by blast + qed + have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" + if "op + 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 "op + 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, 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\Some stepping theorems\ + +lemma dim_empty [simp]: "dim ({} :: 'a::euclidean_space set) = 0" + by (force intro!: dim_unique) + +lemma dim_insert: + fixes x :: "'a::euclidean_space" + shows "dim (insert x S) = (if x \ span S then dim S else dim S + 1)" +proof - + show ?thesis + proof (cases "x \ span S") + case True then show ?thesis + by (metis dim_span span_redundant) + next + case False + obtain B where B: "B \ span S" "independent B" "span S \ span B" "card B = dim (span S)" + using basis_exists [of "span S"] by blast + have 1: "insert x B \ span (insert x S)" + by (meson \B \ span S\ dual_order.trans insertI1 insert_subsetI span_mono span_superset subset_insertI) + have 2: "span (insert x S) \ span (insert x B)" + by (metis \B \ span S\ \span S \ span B\ span_breakdown_eq span_subspace subsetI subspace_span) + have 3: "independent (insert x B)" + by (metis B independent_insert span_subspace subspace_span False) + have "dim (span (insert x S)) = Suc (dim S)" + apply (rule dim_unique [OF 1 2 3]) + by (metis B False card_insert_disjoint dim_span independent_imp_finite subsetCE) + then show ?thesis + by (simp add: False) + qed +qed + +lemma dim_singleton [simp]: + fixes x :: "'a::euclidean_space" + shows "dim{x} = (if x = 0 then 0 else 1)" +by (simp add: dim_insert) + +lemma dim_eq_0 [simp]: + fixes S :: "'a::euclidean_space set" + shows "dim S = 0 \ S \ {0}" +apply safe +apply (metis DIM_positive DIM_real card_ge_dim_independent contra_subsetD dim_empty dim_insert dim_singleton empty_subsetI independent_empty less_not_refl zero_le) +by (metis dim_singleton dim_subset le_0_eq) + +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) + apply (simp add: insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert span_0) + apply (metis (no_types, lifting) add_minus_cancel image_iff uminus_add_conv_diff) + done +qed + +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\General case without assuming closure and getting non-strict separation\ + +proposition 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 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_mul span_superset) + 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_mul span_superset) + 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: divide_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 (op + (- z) ` S)" + by (simp add: \convex S\) + moreover have "op + (- z) ` S \ {}" + by (simp add: \S \ {}\) + moreover have "0 \ op + (- z) ` S" + using zno by auto + ultimately obtain a where "a \ span (op + (- z) ` S)" "a \ 0" + and a: "\x. x \ (op + (- 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 (op + (- z) ` S)\ affine_hull_insert_span_gen apply blast + apply (simp_all add: \a \ 0\ szx) + done +qed + +proposition 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)) (op \ 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\ 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 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 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 + 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 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 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 + +subsection\Misc results about span\ + +lemma eq_span_insert_eq: + assumes "(x - y) \ span S" + shows "span(insert x S) = span(insert y S)" +proof - + have *: "span(insert x S) \ span(insert y S)" if "(x - y) \ span S" for x y + proof - + have 1: "(r *\<^sub>R x - r *\<^sub>R y) \ span S" for r + by (metis real_vector.scale_right_diff_distrib span_mul that) + have 2: "(z - k *\<^sub>R y) - k *\<^sub>R (x - y) = z - k *\<^sub>R x" for z k + by (simp add: real_vector.scale_right_diff_distrib) + show ?thesis + apply (clarsimp simp add: span_breakdown_eq) + by (metis 1 2 diff_add_cancel real_vector.scale_right_diff_distrib span_add_eq) + qed + show ?thesis + apply (intro subset_antisym * assms) + using assms subspace_neg subspace_span minus_diff_eq by force +qed + +lemma dim_psubset: + fixes S :: "'a :: euclidean_space set" + shows "span S \ span T \ dim S < dim T" +by (metis (no_types, hide_lams) dim_span less_le not_le subspace_dim_equal subspace_span) + + +lemma basis_subspace_exists: + fixes S :: "'a::euclidean_space set" + shows + "subspace S + \ \b. finite b \ b \ S \ + independent b \ span b = S \ card b = dim S" +by (metis span_subspace basis_exists independent_imp_finite) + +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_superset) + 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 subs) +qed + +proposition 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 (op + (- c) ` S)" by blast + then have aff: "affine (op + (- c) ` S)" + by (simp add: subspace_imp_affine) + have 0: "0 \ op + (- c) ` S" + by (simp add: \c \ S\) + obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ op + (- c) ` S" + using assms by auto + then have adc: "a \ (d - c) \ 0" + by (simp add: c inner_diff_right) + let ?U = "op + (c+c) ` {x + y |x y. x \ op + (- c) ` S \ a \ y = 0}" + have "u + v \ op + (c + c) ` {x + v |x v. x \ op + (- 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 "... = op + (c+c) ` UNIV" + by (simp add: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) + also have "... = UNIV" + by (simp add: translation_UNIV) + finally show ?thesis . +qed + +proposition dim_sums_Int: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" "subspace T" + shows "dim {x + y |x y. x \ S \ y \ T} + dim(S \ T) = dim S + dim T" +proof - + obtain B where B: "B \ S \ T" "S \ T \ span B" + and indB: "independent B" + and cardB: "card B = dim (S \ T)" + using basis_exists by blast + then obtain C D where "B \ C" "C \ S" "independent C" "S \ span C" + and "B \ D" "D \ T" "independent D" "T \ span D" + using maximal_independent_subset_extend + by (metis Int_subset_iff \B \ S \ T\ indB) + then have "finite B" "finite C" "finite D" + by (simp_all add: independent_imp_finite indB independent_bound) + have Beq: "B = C \ D" + apply (rule sym) + apply (rule spanning_subset_independent) + using \B \ C\ \B \ D\ apply blast + apply (meson \independent C\ independent_mono inf.cobounded1) + using B \C \ S\ \D \ T\ apply auto + done + then have Deq: "D = B \ (D - C)" + by blast + have CUD: "C \ D \ {x + y |x y. x \ S \ y \ T}" + apply safe + apply (metis add.right_neutral subsetCE \C \ S\ \subspace T\ set_eq_subset span_0 span_minimal) + apply (metis add.left_neutral subsetCE \D \ T\ \subspace S\ set_eq_subset span_0 span_minimal) + done + have "a v = 0" if 0: "(\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) = 0" + and v: "v \ C \ (D-C)" for a v + proof - + have eq: "(\v\D - C. a v *\<^sub>R v) = - (\v\C. a v *\<^sub>R v)" + using that add_eq_0_iff by blast + have "(\v\D - C. a v *\<^sub>R v) \ S" + apply (subst eq) + apply (rule subspace_neg [OF \subspace S\]) + apply (rule subspace_sum [OF \subspace S\]) + by (meson subsetCE subspace_mul \C \ S\ \subspace S\) + moreover have "(\v\D - C. a v *\<^sub>R v) \ T" + apply (rule subspace_sum [OF \subspace T\]) + by (meson DiffD1 \D \ T\ \subspace T\ subset_eq subspace_def) + ultimately have "(\v \ D-C. a v *\<^sub>R v) \ span B" + using B by blast + then obtain e where e: "(\v\B. e v *\<^sub>R v) = (\v \ D-C. a v *\<^sub>R v)" + using span_finite [OF \finite B\] by blast + have "\c v. \(\v\C. c v *\<^sub>R v) = 0; v \ C\ \ c v = 0" + using independent_explicit \independent C\ by blast + define cc where "cc x = (if x \ B then a x + e x else a x)" for x + have [simp]: "C \ B = B" "D \ B = B" "C \ - B = C-D" "B \ (D - C) = {}" + using \B \ C\ \B \ D\ Beq by blast+ + have f2: "(\v\C \ D. e v *\<^sub>R v) = (\v\D - C. a v *\<^sub>R v)" + using Beq e by presburger + have f3: "(\v\C \ D. a v *\<^sub>R v) = (\v\C - D. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v) + (\v\C \ D. a v *\<^sub>R v)" + using \finite C\ \finite D\ sum.union_diff2 by blast + have f4: "(\v\C \ (D - C). a v *\<^sub>R v) = (\v\C. a v *\<^sub>R v) + (\v\D - C. a v *\<^sub>R v)" + by (meson Diff_disjoint \finite C\ \finite D\ finite_Diff sum.union_disjoint) + have "(\v\C. cc v *\<^sub>R v) = 0" + using 0 f2 f3 f4 + apply (simp add: cc_def Beq if_smult \finite C\ sum.If_cases algebra_simps sum.distrib) + apply (simp add: add.commute add.left_commute diff_eq) + done + then have "\v. v \ C \ cc v = 0" + using independent_explicit \independent C\ by blast + then have C0: "\v. v \ C - B \ a v = 0" + by (simp add: cc_def Beq) meson + then have [simp]: "(\x\C - B. a x *\<^sub>R x) = 0" + by simp + have "(\x\C. a x *\<^sub>R x) = (\x\B. a x *\<^sub>R x)" + proof - + have "C - D = C - B" + using Beq by blast + then show ?thesis + using Beq \(\x\C - B. a x *\<^sub>R x) = 0\ f3 f4 by auto + qed + with 0 have Dcc0: "(\v\D. a v *\<^sub>R v) = 0" + apply (subst Deq) + by (simp add: \finite B\ \finite D\ sum_Un) + then have D0: "\v. v \ D \ a v = 0" + using independent_explicit \independent D\ by blast + show ?thesis + using v C0 D0 Beq by blast + qed + then have "independent (C \ (D - C))" + by (simp add: independent_explicit \finite C\ \finite D\ sum_Un del: Un_Diff_cancel) + then have indCUD: "independent (C \ D)" by simp + have "dim (S \ T) = card B" + by (rule dim_unique [OF B indB refl]) + moreover have "dim S = card C" + by (metis \C \ S\ \independent C\ \S \ span C\ basis_card_eq_dim) + moreover have "dim T = card D" + by (metis \D \ T\ \independent D\ \T \ span D\ basis_card_eq_dim) + moreover have "dim {x + y |x y. x \ S \ y \ T} = card(C \ D)" + apply (rule dim_unique [OF CUD _ indCUD refl], clarify) + apply (meson \S \ span C\ \T \ span D\ span_add span_inc span_minimal subsetCE subspace_span sup.bounded_iff) + done + ultimately show ?thesis + using \B = C \ D\ [symmetric] + by (simp add: \independent C\ \independent D\ card_Un_Int independent_finite) +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 (op+ (-a) ` S)" "affine (op+ (-a) ` T)" + using assms by (auto simp: affine_translation [symmetric]) + have zero: "0 \ (op+ (-a) ` S)" "0 \ (op+ (-a) ` T)" + using a assms by auto + have [simp]: "{x + y |x y. x \ op + (- a) ` S \ y \ op + (- a) ` T} = + op + (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" + by (force simp: algebra_simps scaleR_2) + have [simp]: "op + (- a) ` S \ op + (- a) ` T = op + (- a) ` (S \ T)" + by auto + show ?thesis + using aff_dim_sums_Int_0 [OF aff zero] + by (auto simp: aff_dim_translation_eq) +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 dim_Times: + fixes S :: "'a :: euclidean_space set" and T :: "'a set" + assumes "subspace S" "subspace T" + shows "dim(S \ T) = dim S + dim T" +proof - + have ss: "subspace ((\x. (x, 0)) ` S)" "subspace (Pair 0 ` T)" + by (rule subspace_linear_image, unfold_locales, auto simp: assms)+ + have "dim(S \ T) = dim({u + v |u v. u \ (\x. (x, 0)) ` S \ v \ Pair 0 ` T})" + by (simp add: Times_eq_image_sum) + moreover have "dim ((\x. (x, 0::'a)) ` S) = dim S" "dim (Pair (0::'a) ` T) = dim T" + by (auto simp: additive.intro linear.intro linear_axioms.intro inj_on_def intro: dim_image_eq) + moreover have "dim ((\x. (x, 0)) ` S \ Pair 0 ` T) = 0" + by (subst dim_eq_0) (force simp: zero_prod_def) + ultimately show ?thesis + using dim_sums_Int [OF ss] by linarith +qed + +subsection\ Orthogonal bases, Gram-Schmidt process, and related theorems\ + +lemma span_delete_0 [simp]: "span(S - {0}) = span S" +proof + show "span (S - {0}) \ span S" + by (blast intro!: span_mono) +next + have "span S \ span(insert 0 (S - {0}))" + by (blast intro!: span_mono) + also have "... \ span(S - {0})" + using span_insert_0 by blast + finally show "span S \ span (S - {0})" . +qed + +lemma span_image_scale: + assumes "finite S" and nz: "\x. x \ S \ c x \ 0" + shows "span ((\x. c x *\<^sub>R x) ` S) = span S" +using assms +proof (induction S arbitrary: c) + case (empty c) show ?case by simp +next + case (insert x F c) + show ?case + proof (intro set_eqI iffI) + fix y + assume "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" + then show "y \ span (insert x F)" + using insert by (force simp: span_breakdown_eq) + next + fix y + assume "y \ span (insert x F)" + then show "y \ span ((\x. c x *\<^sub>R x) ` insert x F)" + using insert + apply (clarsimp simp: span_breakdown_eq) + apply (rule_tac x="k / c x" in exI) + by simp + qed +qed + +lemma pairwise_orthogonal_independent: + assumes "pairwise orthogonal S" and "0 \ S" + shows "independent S" +proof - + have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" + using assms by (simp add: pairwise_def orthogonal_def) + have "False" if "a \ S" and a: "a \ span (S - {a})" for a + proof - + obtain T U where "T \ S - {a}" "a = (\v\T. U v *\<^sub>R v)" + using a by (force simp: span_explicit) + then have "a \ a = a \ (\v\T. U v *\<^sub>R v)" + by simp + also have "... = 0" + apply (simp add: inner_sum_right) + apply (rule comm_monoid_add_class.sum.neutral) + by (metis "0" DiffE \T \ S - {a}\ mult_not_zero singletonI subsetCE \a \ S\) + finally show ?thesis + using \0 \ S\ \a \ S\ by auto + qed + then show ?thesis + by (force simp: dependent_def) +qed + +lemma pairwise_orthogonal_imp_finite: + fixes S :: "'a::euclidean_space set" + assumes "pairwise orthogonal S" + shows "finite S" +proof - + have "independent (S - {0})" + apply (rule pairwise_orthogonal_independent) + apply (metis Diff_iff assms pairwise_def) + by blast + then show ?thesis + by (meson independent_imp_finite infinite_remove) +qed + +lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}" + by (simp add: subspace_def orthogonal_clauses) + +lemma subspace_orthogonal_to_vectors: "subspace {y. \x \ S. orthogonal x y}" + by (simp add: subspace_def orthogonal_clauses) + +lemma orthogonal_to_span: + assumes a: "a \ span S" and x: "\y. y \ S \ orthogonal x y" + shows "orthogonal x a" +apply (rule span_induct [OF a subspace_orthogonal_to_vector]) +apply (simp add: x) +done + +proposition Gram_Schmidt_step: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" and x: "x \ span S" + shows "orthogonal x (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b))" +proof - + have "finite S" + by (simp add: S pairwise_orthogonal_imp_finite) + have "orthogonal (a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)) x" + if "x \ S" for x + proof - + have "a \ x = (\y\S. if y = x then y \ a else 0)" + by (simp add: \finite S\ inner_commute sum.delta that) + also have "... = (\b\S. b \ a * (b \ x) / (b \ b))" + apply (rule sum.cong [OF refl], simp) + by (meson S orthogonal_def pairwise_def that) + finally show ?thesis + by (simp add: orthogonal_def algebra_simps inner_sum_left) + qed + then show ?thesis + using orthogonal_to_span orthogonal_commute x by blast +qed + + +lemma orthogonal_extension_aux: + fixes S :: "'a::euclidean_space set" + assumes "finite T" "finite S" "pairwise orthogonal S" + shows "\U. pairwise orthogonal (S \ U) \ span (S \ U) = span (S \ T)" +using assms +proof (induction arbitrary: S) + case empty then show ?case + by simp (metis sup_bot_right) +next + case (insert a T) + have 0: "\x y. \x \ y; x \ S; y \ S\ \ x \ y = 0" + using insert by (simp add: pairwise_def orthogonal_def) + define a' where "a' = a - (\b\S. (b \ a / (b \ b)) *\<^sub>R b)" + obtain U where orthU: "pairwise orthogonal (S \ insert a' U)" + and spanU: "span (insert a' S \ U) = span (insert a' S \ T)" + apply (rule exE [OF insert.IH [of "insert a' S"]]) + apply (auto simp: Gram_Schmidt_step a'_def insert.prems orthogonal_commute pairwise_orthogonal_insert span_clauses) + done + have orthS: "\x. x \ S \ a' \ x = 0" + apply (simp add: a'_def) + using Gram_Schmidt_step [OF \pairwise orthogonal S\] + apply (force simp: orthogonal_def inner_commute span_inc [THEN subsetD]) + done + have "span (S \ insert a' U) = span (insert a' (S \ T))" + using spanU by simp + also have "... = span (insert a (S \ T))" + apply (rule eq_span_insert_eq) + apply (simp add: a'_def span_neg span_sum span_clauses(1) span_mul) + done + also have "... = span (S \ insert a T)" + by simp + finally show ?case + apply (rule_tac x="insert a' U" in exI) + using orthU apply auto + done +qed + + +proposition orthogonal_extension: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" + obtains U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" +proof - + obtain B where "finite B" "span B = span T" + using basis_subspace_exists [of "span T"] subspace_span by auto + with orthogonal_extension_aux [of B S] + obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ B)" + using assms pairwise_orthogonal_imp_finite by auto + show ?thesis + apply (rule_tac U=U in that) + apply (simp add: \pairwise orthogonal (S \ U)\) + by (metis \span (S \ U) = span (S \ B)\ \span B = span T\ span_Un) +qed + +corollary orthogonal_extension_strong: + fixes S :: "'a::euclidean_space set" + assumes S: "pairwise orthogonal S" + obtains U where "U \ (insert 0 S) = {}" "pairwise orthogonal (S \ U)" + "span (S \ U) = span (S \ T)" +proof - + obtain U where "pairwise orthogonal (S \ U)" "span (S \ U) = span (S \ T)" + using orthogonal_extension assms by blast + then show ?thesis + apply (rule_tac U = "U - (insert 0 S)" in that) + apply blast + apply (force simp: pairwise_def) + apply (metis (no_types, lifting) Un_Diff_cancel span_insert_0 span_Un) + done +qed + +subsection\Decomposing a vector into parts in orthogonal subspaces.\ + +text\existence of orthonormal basis for a subspace.\ + +lemma orthogonal_spanningset_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "B \ S" "pairwise orthogonal B" "span B = S" +proof - + obtain B where "B \ S" "independent B" "S \ span B" "card B = dim S" + using basis_exists by blast + with orthogonal_extension [of "{}" B] + show ?thesis + by (metis Un_empty_left assms pairwise_empty span_inc span_subspace that) +qed + +lemma orthogonal_basis_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "0 \ B" "B \ S" "pairwise orthogonal B" "independent B" + "card B = dim S" "span B = S" +proof - + obtain B where "B \ S" "pairwise orthogonal B" "span B = S" + using assms orthogonal_spanningset_subspace by blast + then show ?thesis + apply (rule_tac B = "B - {0}" in that) + apply (auto simp: indep_card_eq_dim_span pairwise_subset Diff_subset pairwise_orthogonal_independent elim: pairwise_subset) + done +qed + +proposition orthonormal_basis_subspace: + fixes S :: "'a :: euclidean_space set" + assumes "subspace S" + obtains B where "B \ S" "pairwise orthogonal B" + and "\x. x \ B \ norm x = 1" + and "independent B" "card B = dim S" "span B = S" +proof - + obtain B where "0 \ B" "B \ S" + and orth: "pairwise orthogonal B" + and "independent B" "card B = dim S" "span B = S" + by (blast intro: orthogonal_basis_subspace [OF assms]) + have 1: "(\x. x /\<^sub>R norm x) ` B \ S" + using \span B = S\ span_clauses(1) span_mul by fastforce + have 2: "pairwise orthogonal ((\x. x /\<^sub>R norm x) ` B)" + using orth by (force simp: pairwise_def orthogonal_clauses) + have 3: "\x. x \ (\x. x /\<^sub>R norm x) ` B \ norm x = 1" + by (metis (no_types, lifting) \0 \ B\ image_iff norm_sgn sgn_div_norm) + have 4: "independent ((\x. x /\<^sub>R norm x) ` B)" + by (metis "2" "3" norm_zero pairwise_orthogonal_independent zero_neq_one) + have "inj_on (\x. x /\<^sub>R norm x) B" + proof + fix x y + assume "x \ B" "y \ B" "x /\<^sub>R norm x = y /\<^sub>R norm y" + moreover have "\i. i \ B \ norm (i /\<^sub>R norm i) = 1" + using 3 by blast + ultimately show "x = y" + by (metis norm_eq_1 orth orthogonal_clauses(7) orthogonal_commute orthogonal_def pairwise_def zero_neq_one) + qed + then have 5: "card ((\x. x /\<^sub>R norm x) ` B) = dim S" + by (metis \card B = dim S\ card_image) + have 6: "span ((\x. x /\<^sub>R norm x) ` B) = S" + by (metis "1" "4" "5" assms card_eq_dim independent_finite span_subspace) + show ?thesis + by (rule that [OF 1 2 3 4 5 6]) +qed + +proposition orthogonal_subspace_decomp_exists: + fixes S :: "'a :: euclidean_space set" + obtains y z where "y \ span S" "\w. w \ span S \ orthogonal z w" "x = y + z" +proof - + obtain T where "0 \ T" "T \ span S" "pairwise orthogonal T" "independent T" "card T = dim (span S)" "span T = span S" + using orthogonal_basis_subspace subspace_span by blast + let ?a = "\b\T. (b \ x / (b \ b)) *\<^sub>R b" + have orth: "orthogonal (x - ?a) w" if "w \ span S" for w + by (simp add: Gram_Schmidt_step \pairwise orthogonal T\ \span T = span S\ orthogonal_commute that) + show ?thesis + apply (rule_tac y = "?a" and z = "x - ?a" in that) + apply (meson \T \ span S\ span_mul span_sum subsetCE) + apply (fact orth, simp) + done +qed + +lemma orthogonal_subspace_decomp_unique: + fixes S :: "'a :: euclidean_space set" + assumes "x + y = x' + y'" + and ST: "x \ span S" "x' \ span S" "y \ span T" "y' \ span T" + and orth: "\a b. \a \ S; b \ T\ \ orthogonal a b" + shows "x = x' \ y = y'" +proof - + have "x + y - y' = x'" + by (simp add: assms) + moreover have "\a b. \a \ span S; b \ span T\ \ orthogonal a b" + by (meson orth orthogonal_commute orthogonal_to_span) + ultimately have "0 = x' - x" + by (metis (full_types) add_diff_cancel_left' ST diff_right_commute orthogonal_clauses(10) orthogonal_clauses(5) orthogonal_self) + with assms show ?thesis by auto +qed + +proposition dim_orthogonal_sum: + fixes A :: "'a::euclidean_space set" + assumes "\x y. \x \ A; y \ B\ \ x \ y = 0" + shows "dim(A \ B) = dim A + dim B" +proof - + have 1: "\x y. \x \ span A; y \ B\ \ x \ y = 0" + by (erule span_induct [OF _ subspace_hyperplane2]; simp add: assms) + have "\x y. \x \ span A; y \ span B\ \ x \ y = 0" + apply (erule span_induct [OF _ subspace_hyperplane]) + using 1 by (simp add: ) + then have 0: "\x y. \x \ span A; y \ span B\ \ x \ y = 0" + by simp + have "dim(A \ B) = dim (span (A \ B))" + by simp + also have "... = dim ((\(a, b). a + b) ` (span A \ span B))" + by (simp add: span_Un) + also have "... = dim {x + y |x y. x \ span A \ y \ span B}" + by (auto intro!: arg_cong [where f=dim]) + also have "... = dim {x + y |x y. x \ span A \ y \ span B} + dim(span A \ span B)" + by (auto simp: dest: 0) + also have "... = dim (span A) + dim (span B)" + by (rule dim_sums_Int) auto + also have "... = dim A + dim B" + by simp + finally show ?thesis . +qed + +lemma dim_subspace_orthogonal_to_vectors: + fixes A :: "'a::euclidean_space set" + assumes "subspace A" "subspace B" "A \ B" + shows "dim {y \ B. \x \ A. orthogonal x y} + dim A = dim B" +proof - + have "dim (span ({y \ B. \x\A. orthogonal x y} \ A)) = dim (span B)" + proof (rule arg_cong [where f=dim, OF subset_antisym]) + show "span ({y \ B. \x\A. orthogonal x y} \ A) \ span B" + by (simp add: \A \ B\ Collect_restrict span_mono) + next + have *: "x \ span ({y \ B. \x\A. orthogonal x y} \ A)" + if "x \ B" for x + proof - + obtain y z where "x = y + z" "y \ span A" and orth: "\w. w \ span A \ orthogonal z w" + using orthogonal_subspace_decomp_exists [of A x] that by auto + have "y \ span B" + by (metis span_eq \y \ span A\ assms subset_iff) + then have "z \ {a \ B. \x. x \ A \ orthogonal x a}" + by simp (metis (no_types) span_eq \x = y + z\ \subspace A\ \subspace B\ orth orthogonal_commute span_add_eq that) + then have z: "z \ span {y \ B. \x\A. orthogonal x y}" + by (meson span_inc subset_iff) + then show ?thesis + apply (simp add: span_Un image_def) + apply (rule bexI [OF _ z]) + apply (simp add: \x = y + z\ \y \ span A\) + done + qed + show "span B \ span ({y \ B. \x\A. orthogonal x y} \ A)" + by (rule span_minimal) (auto intro: * span_minimal elim: ) + qed + then show ?thesis + by (metis (no_types, lifting) dim_orthogonal_sum dim_span mem_Collect_eq orthogonal_commute orthogonal_def) +qed + +lemma aff_dim_openin: + fixes S :: "'a::euclidean_space set" + assumes ope: "openin (subtopology euclidean 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 \independent B\ + apply (rule independent_injective_image, simp) + by (metis \0 < e\ injective_scaleR less_irrefl) + qed + also have "... = aff_dim S" + using \a \ S\ aff_dim_eq_dim hull_inc by force + finally show "aff_dim T \ aff_dim S" . + qed +qed + +lemma dim_openin: + fixes S :: "'a::euclidean_space set" + assumes ope: "openin (subtopology euclidean 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_inc subspace_span) + finally show "dim T \ dim S" by simp +qed + +subsection\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 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 (op + (- z) ` S)" by blast + then have aff: "affine (op + (- z) ` S)" and span: "span (op + (- z) ` S) = (op + (- z) ` S)" + by (auto simp: subspace_imp_affine) + obtain a' a'' where a': "a' \ span (op + (- z) ` S)" and a: "a = a' + a''" + and "\w. w \ span (op + (- z) ` S) \ orthogonal a'' w" + using orthogonal_subspace_decomp_exists [of "op + (- z) ` S" "a"] by metis + then have "\w. w \ S \ a'' \ (w-z) = 0" + by (simp add: imageI orthogonal_def span) + 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 \ op + (- z) ` S \ (w + a') \ op + (- z) ` S" + by (metis subspace_add a' span_eq 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\) + 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) + 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 linear_injective_left_inverse assms 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 + + +text\Choosing a subspace of a given dimension\ +proposition choose_subspace_of_subspace: + fixes S :: "'n::euclidean_space set" + assumes "n \ dim S" + obtains T where "subspace T" "T \ span S" "dim T = n" +proof - + have "\T. subspace T \ T \ span S \ dim T = n" + using assms + proof (induction n) + case 0 then show ?case by force + next + case (Suc n) + then obtain T where "subspace T" "T \ span S" "dim T = n" + by force + then show ?case + proof (cases "span S \ span T") + case True + have "dim S = dim T" + apply (rule span_eq_dim [OF subset_antisym [OF True]]) + by (simp add: \T \ span S\ span_minimal subspace_span) + then show ?thesis + using Suc.prems \dim T = n\ by linarith + next + case False + then obtain y where y: "y \ S" "y \ T" + by (meson span_mono subsetI) + then have "span (insert y T) \ span S" + by (metis (no_types) \T \ span S\ subsetD insert_subset span_inc span_mono span_span) + with \dim T = n\ \subspace T\ y show ?thesis + apply (rule_tac x="span(insert y T)" in exI) + apply (auto simp: dim_insert) + using span_eq by blast + qed + qed + with that show ?thesis by blast +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 (op + (- a) ` S)" + by (simp add: affine_diffs_subspace) + have "nat d \ dim (op + (- 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 (op + (- a) ` S)" + and Tdim: "dim T = nat d" + using choose_subspace_of_subspace [of "nat d" "op + (- a) ` S"] by blast + then have "affine T" + using subspace_affine by blast + then have "affine (op + a ` T)" + by (metis affine_hull_eq affine_hull_translation) + moreover have "op + a ` T \ S" + proof - + have "T \ op + (- a) ` S" + by (metis (no_types) span_eq Tsb ss) + then show "op + a ` T \ S" + using add_ac by auto + qed + moreover have "aff_dim (op + 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\Several Variants of Paracompactness\ + +proposition paracompact: + fixes S :: "'a :: euclidean_space 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) + done + 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 \" "\\ = UNION S G" + using Lindelof [of "G ` S"] by (metis image_iff) + then obtain K where K: "K \ S" "countable K" and eq: "UNION K G = UNION S G" + 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 :: euclidean_space set" + assumes cin: "closedin (subtopology euclidean U) S" + and oin: "\T. T \ \ \ openin (subtopology euclidean U) T" + and "S \ \\" + obtains \' where "S \ \ \'" + and "\V. V \ \' \ openin (subtopology euclidean U) V \ (\T. T \ \ \ V \ T)" + and "\x. x \ U + \ \V. openin (subtopology euclidean 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 \ {}}" + using paracompact [OF 1 2] by auto + 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 (subtopology euclidean 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 (subtopology euclidean 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 paracompact_closed: + fixes S :: "'a :: euclidean_space 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 \ {})}" +using paracompact_closedin [of UNIV S \] assms +by (auto simp: open_openin [symmetric] closed_closedin [symmetric]) + + +subsection\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 (subtopology euclidean (S \ T)) ((\x. Pair x (f x)) ` S)" +proof - + have eq: "((\x. Pair x (f x)) ` S) = {z. z \ S \ T \ (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 (subtopology euclidean (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 (subtopology euclidean T) U" + have eq: "{x. x \ S \ f x \ U} = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" + by (force simp: image_iff) + show "closedin (subtopology euclidean S) {x \ S. f x \ 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 (subtopology euclidean (S \ T)) S" + and opT: "openin (subtopology euclidean (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}" "S \ T" "\i. i" "\i. f" f] contf contg opS opT by blast + +lemma continuous_on_cases_local_open: + assumes opS: "openin (subtopology euclidean (S \ T)) S" + and opT: "openin (subtopology euclidean (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 + +subsection\The union of two collinear segments is another segment\ + +proposition 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 real_vector.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 real_vector.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\ \UNION UNIV f = 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 divide_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 \ UNION UNIV ?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: divide_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 \ UNION UNIV ?C" + by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) + qed + then show "UNION UNIV ?C = S" by auto + qed + ultimately show ?thesis + using that by metis +qed + +end + \ No newline at end of file