diff --git a/src/HOL/Analysis/Convex.thy b/src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy +++ b/src/HOL/Analysis/Convex.thy @@ -1,2493 +1,2453 @@ (* Title: HOL/Analysis/Convex.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Convex Sets and Functions\ theory Convex imports Affine "HOL-Library.Set_Algebras" begin subsection \Convex Sets\ definition\<^marker>\tag important\ convex :: "'a::real_vector set \ bool" where "convex s \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" lemma convexI: assumes "\x y u v. x \ s \ y \ s \ 0 \ u \ 0 \ v \ u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" shows "convex s" using assms unfolding convex_def by fast lemma convexD: assumes "convex s" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u *\<^sub>R x + v *\<^sub>R y \ s" using assms unfolding convex_def by fast lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R y) \ s)" (is "_ \ ?alt") proof show "convex s" if alt: ?alt proof - { fix x y and u v :: real assume mem: "x \ s" "y \ s" assume "0 \ u" "0 \ v" moreover assume "u + v = 1" then have "u = 1 - v" by auto ultimately have "u *\<^sub>R x + v *\<^sub>R y \ s" using alt [rule_format, OF mem] by auto } then show ?thesis unfolding convex_def by auto qed show ?alt if "convex s" using that by (auto simp: convex_def) qed lemma convexD_alt: assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \ s" using assms unfolding convex_alt by auto lemma mem_convex_alt: assumes "convex S" "x \ S" "y \ S" "u \ 0" "v \ 0" "u + v > 0" shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \ S" apply (rule convexD) using assms apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) done lemma convex_empty[intro,simp]: "convex {}" unfolding convex_def by simp lemma convex_singleton[intro,simp]: "convex {a}" unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) lemma convex_UNIV[intro,simp]: "convex UNIV" unfolding convex_def by auto lemma convex_Inter: "(\s. s\f \ convex s) \ convex(\f)" unfolding convex_def by auto lemma convex_Int: "convex s \ convex t \ convex (s \ t)" unfolding convex_def by auto lemma convex_INT: "(\i. i \ A \ convex (B i)) \ convex (\i\A. B i)" unfolding convex_def by auto lemma convex_Times: "convex s \ convex t \ convex (s \ t)" unfolding convex_def by auto lemma convex_halfspace_le: "convex {x. inner a x \ b}" unfolding convex_def by (auto simp: inner_add intro!: convex_bound_le) lemma convex_halfspace_ge: "convex {x. inner a x \ b}" proof - have *: "{x. inner a x \ b} = {x. inner (-a) x \ -b}" by auto show ?thesis unfolding * using convex_halfspace_le[of "-a" "-b"] by auto qed lemma convex_halfspace_abs_le: "convex {x. \inner a x\ \ b}" proof - have *: "{x. \inner a x\ \ b} = {x. inner a x \ b} \ {x. -b \ inner a x}" by auto show ?thesis unfolding * by (simp add: convex_Int convex_halfspace_ge convex_halfspace_le) qed lemma convex_hyperplane: "convex {x. inner a x = b}" proof - have *: "{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" by auto show ?thesis using convex_halfspace_le convex_halfspace_ge by (auto intro!: convex_Int simp: *) qed lemma convex_halfspace_lt: "convex {x. inner a x < b}" unfolding convex_def by (auto simp: convex_bound_lt inner_add) lemma convex_halfspace_gt: "convex {x. inner a x > b}" using convex_halfspace_lt[of "-a" "-b"] by auto lemma convex_halfspace_Re_ge: "convex {x. Re x \ b}" using convex_halfspace_ge[of b "1::complex"] by simp lemma convex_halfspace_Re_le: "convex {x. Re x \ b}" using convex_halfspace_le[of "1::complex" b] by simp lemma convex_halfspace_Im_ge: "convex {x. Im x \ b}" using convex_halfspace_ge[of b \] by simp lemma convex_halfspace_Im_le: "convex {x. Im x \ b}" using convex_halfspace_le[of \ b] by simp lemma convex_halfspace_Re_gt: "convex {x. Re x > b}" using convex_halfspace_gt[of b "1::complex"] by simp lemma convex_halfspace_Re_lt: "convex {x. Re x < b}" using convex_halfspace_lt[of "1::complex" b] by simp lemma convex_halfspace_Im_gt: "convex {x. Im x > b}" using convex_halfspace_gt[of b \] by simp lemma convex_halfspace_Im_lt: "convex {x. Im x < b}" using convex_halfspace_lt[of \ b] by simp lemma convex_real_interval [iff]: fixes a b :: "real" shows "convex {a..}" and "convex {..b}" and "convex {a<..}" and "convex {.. inner 1 x}" by auto then show 1: "convex {a..}" by (simp only: convex_halfspace_ge) have "{..b} = {x. inner 1 x \ b}" by auto then show 2: "convex {..b}" by (simp only: convex_halfspace_le) have "{a<..} = {x. a < inner 1 x}" by auto then show 3: "convex {a<..}" by (simp only: convex_halfspace_gt) have "{.. {..b}" by auto then show "convex {a..b}" by (simp only: convex_Int 1 2) have "{a<..b} = {a<..} \ {..b}" by auto then show "convex {a<..b}" by (simp only: convex_Int 3 2) have "{a.. {.. {.." by (simp add: convex_def scaleR_conv_of_real) subsection\<^marker>\tag unimportant\ \Explicit expressions for convexity in terms of arbitrary sums\ lemma convex_sum: fixes C :: "'a::real_vector set" assumes "finite s" and "convex C" and "(\ i \ s. a i) = 1" assumes "\i. i \ s \ a i \ 0" and "\i. i \ s \ y i \ C" shows "(\ j \ s. a j *\<^sub>R y j) \ C" using assms(1,3,4,5) proof (induct arbitrary: a set: finite) case empty then show ?case by simp next case (insert i s) note IH = this(3) have "a i + sum a s = 1" and "0 \ a i" and "\j\s. 0 \ a j" and "y i \ C" and "\j\s. y j \ C" using insert.hyps(1,2) insert.prems by simp_all then have "0 \ sum a s" by (simp add: sum_nonneg) have "a i *\<^sub>R y i + (\j\s. a j *\<^sub>R y j) \ C" proof (cases "sum a s = 0") case True with \a i + sum a s = 1\ have "a i = 1" by simp from sum_nonneg_0 [OF \finite s\ _ True] \\j\s. 0 \ a j\ have "\j\s. a j = 0" by simp show ?thesis using \a i = 1\ and \\j\s. a j = 0\ and \y i \ C\ by simp next case False with \0 \ sum a s\ have "0 < sum a s" by simp then have "(\j\s. (a j / sum a s) *\<^sub>R y j) \ C" using \\j\s. 0 \ a j\ and \\j\s. y j \ C\ by (simp add: IH sum_divide_distrib [symmetric]) from \convex C\ and \y i \ C\ and this and \0 \ a i\ and \0 \ sum a s\ and \a i + sum a s = 1\ have "a i *\<^sub>R y i + sum a s *\<^sub>R (\j\s. (a j / sum a s) *\<^sub>R y j) \ C" by (rule convexD) then show ?thesis by (simp add: scaleR_sum_right False) qed then show ?case using \finite s\ and \i \ s\ by simp qed lemma convex: "convex s \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (sum u {1..k} = 1) \ sum (\i. u i *\<^sub>R x i) {1..k} \ s)" proof safe fix k :: nat fix u :: "nat \ real" fix x assume "convex s" "\i. 1 \ i \ i \ k \ 0 \ u i \ x i \ s" "sum u {1..k} = 1" with convex_sum[of "{1 .. k}" s] show "(\j\{1 .. k}. u j *\<^sub>R x j) \ s" by auto next assume *: "\k u x. (\ i :: nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ sum u {1..k} = 1 \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ s" { fix \ :: real fix x y :: 'a assume xy: "x \ s" "y \ s" assume mu: "\ \ 0" "\ \ 1" let ?u = "\i. if (i :: nat) = 1 then \ else 1 - \" let ?x = "\i. if (i :: nat) = 1 then x else y" have "{1 :: nat .. 2} \ - {x. x = 1} = {2}" by auto then have card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" by simp then have "sum ?u {1 .. 2} = 1" using sum.If_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] by auto with *[rule_format, of "2" ?u ?x] have s: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ s" using mu xy by auto have grarr: "(\j \ {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \) *\<^sub>R y" using sum.atLeast_Suc_atMost[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto from sum.atLeast_Suc_atMost[of "Suc 0" 2 "\ j. ?u j *\<^sub>R ?x j", simplified this] have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" by auto then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" using s by (auto simp: add.commute) } then show "convex s" unfolding convex_alt by auto qed lemma convex_explicit: fixes s :: "'a::real_vector set" shows "convex s \ (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ sum u t = 1 \ sum (\x. u x *\<^sub>R x) t \ s)" proof safe fix t fix u :: "'a \ real" assume "convex s" and "finite t" and "t \ s" "\x\t. 0 \ u x" "sum u t = 1" then show "(\x\t. u x *\<^sub>R x) \ s" using convex_sum[of t s u "\ x. x"] by auto next assume *: "\t. \ u. finite t \ t \ s \ (\x\t. 0 \ u x) \ sum u t = 1 \ (\x\t. u x *\<^sub>R x) \ s" show "convex s" unfolding convex_alt proof safe fix x y fix \ :: real assume **: "x \ s" "y \ s" "0 \ \" "\ \ 1" show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" proof (cases "x = y") case False then show ?thesis using *[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] ** by auto next case True then show ?thesis using *[rule_format, of "{x, y}" "\ z. 1"] ** by (auto simp: field_simps real_vector.scale_left_diff_distrib) qed qed qed lemma convex_finite: assumes "finite s" shows "convex s \ (\u. (\x\s. 0 \ u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s \ s)" unfolding convex_explicit apply safe subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto subgoal for t u proof - have if_distrib_arg: "\P f g x. (if P then f else g) x = (if P then f x else g x)" by simp assume sum: "\u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) \ s" assume *: "\x\t. 0 \ u x" "sum u t = 1" assume "t \ s" then have "s \ t = t" by auto with sum[THEN spec[where x="\x. if x\t then u x else 0"]] * show "(\x\t. u x *\<^sub>R x) \ s" by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) qed done subsection \Convex Functions on a Set\ definition\<^marker>\tag important\ convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where "convex_on s f \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y)" lemma convex_onI [intro?]: assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" shows "convex_on A f" unfolding convex_on_def proof clarify fix x y fix u v :: real assume A: "x \ A" "y \ A" "u \ 0" "v \ 0" "u + v = 1" from A(5) have [simp]: "v = 1 - u" by (simp add: algebra_simps) from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" using assms[of u y x] by (cases "u = 0 \ u = 1") (auto simp: algebra_simps) qed lemma convex_on_linorderI [intro?]: fixes A :: "('a::{linorder,real_vector}) set" assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ x < y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" shows "convex_on A f" proof fix x y fix t :: real assume A: "x \ A" "y \ A" "t > 0" "t < 1" with assms [of t x y] assms [of "1 - t" y x] show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" by (cases x y rule: linorder_cases) (auto simp: algebra_simps) qed lemma convex_onD: assumes "convex_on A f" shows "\t x y. t \ 0 \ t \ 1 \ x \ A \ y \ A \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" using assms by (auto simp: convex_on_def) lemma convex_onD_Icc: assumes "convex_on {x..y} f" "x \ (y :: _ :: {real_vector,preorder})" shows "\t. t \ 0 \ t \ 1 \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" using assms(2) by (intro convex_onD [OF assms(1)]) simp_all lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" unfolding convex_on_def by auto lemma convex_on_add [intro]: assumes "convex_on s f" and "convex_on s g" shows "convex_on s (\x. f x + g x)" proof - { fix x y assume "x \ s" "y \ s" moreover fix u v :: real assume "0 \ u" "0 \ v" "u + v = 1" ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" using assms unfolding convex_on_def by (auto simp: add_mono) then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps) } then show ?thesis unfolding convex_on_def by auto qed lemma convex_on_cmul [intro]: fixes c :: real assumes "0 \ c" and "convex_on s f" shows "convex_on s (\x. c * f x)" proof - have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" for u c fx v fy :: real by (simp add: field_simps) show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] unfolding convex_on_def and * by auto qed lemma convex_lower: assumes "convex_on s f" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" shows "f (u *\<^sub>R x + v *\<^sub>R y) \ max (f x) (f y)" proof - let ?m = "max (f x) (f y)" have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" using assms(4,5) by (auto simp: mult_left_mono add_mono) also have "\ = max (f x) (f y)" using assms(6) by (simp add: distrib_right [symmetric]) finally show ?thesis using assms unfolding convex_on_def by fastforce qed lemma convex_on_dist [intro]: fixes s :: "'a::real_normed_vector set" shows "convex_on s (\x. dist a x)" proof (auto simp: convex_on_def dist_norm) fix x y assume "x \ s" "y \ s" fix u v :: real assume "0 \ u" assume "0 \ v" assume "u + v = 1" have "a = u *\<^sub>R a + v *\<^sub>R a" unfolding scaleR_left_distrib[symmetric] and \u + v = 1\ by simp then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" by (auto simp: algebra_simps) show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] using \0 \ u\ \0 \ v\ by auto qed subsection\<^marker>\tag unimportant\ \Arithmetic operations on sets preserve convexity\ lemma convex_linear_image: assumes "linear f" and "convex s" shows "convex (f ` s)" proof - interpret f: linear f by fact from \convex s\ show "convex (f ` s)" by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) qed lemma convex_linear_vimage: assumes "linear f" and "convex s" shows "convex (f -` s)" proof - interpret f: linear f by fact from \convex s\ show "convex (f -` s)" by (simp add: convex_def f.add f.scaleR) qed lemma convex_scaling: assumes "convex s" shows "convex ((\x. c *\<^sub>R x) ` s)" proof - have "linear (\x. c *\<^sub>R x)" by (simp add: linearI scaleR_add_right) then show ?thesis using \convex s\ by (rule convex_linear_image) qed lemma convex_scaled: assumes "convex S" shows "convex ((\x. x *\<^sub>R c) ` S)" proof - have "linear (\x. x *\<^sub>R c)" by (simp add: linearI scaleR_add_left) then show ?thesis using \convex S\ by (rule convex_linear_image) qed lemma convex_negations: assumes "convex S" shows "convex ((\x. - x) ` S)" proof - have "linear (\x. - x)" by (simp add: linearI) then show ?thesis using \convex S\ by (rule convex_linear_image) qed lemma convex_sums: assumes "convex S" and "convex T" shows "convex (\x\ S. \y \ T. {x + y})" proof - have "linear (\(x, y). x + y)" by (auto intro: linearI simp: scaleR_add_right) with assms have "convex ((\(x, y). x + y) ` (S \ T))" by (intro convex_linear_image convex_Times) also have "((\(x, y). x + y) ` (S \ T)) = (\x\ S. \y \ T. {x + y})" by auto finally show ?thesis . qed lemma convex_differences: assumes "convex S" "convex T" shows "convex (\x\ S. \y \ T. {x - y})" proof - have "{x - y| x y. x \ S \ y \ T} = {x + y |x y. x \ S \ y \ uminus ` T}" by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff) then show ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto qed lemma convex_translation: "convex ((+) a ` S)" if "convex S" proof - have "(\ x\ {a}. \y \ S. {x + y}) = (+) a ` S" by auto then show ?thesis using convex_sums [OF convex_singleton [of a] that] by auto qed lemma convex_translation_subtract: "convex ((\b. b - a) ` S)" if "convex S" using convex_translation [of S "- a"] that by (simp cong: image_cong_simp) lemma convex_affinity: assumes "convex S" shows "convex ((\x. a + c *\<^sub>R x) ` S)" proof - have "(\x. a + c *\<^sub>R x) ` S = (+) a ` (*\<^sub>R) c ` S" by auto then show ?thesis using convex_translation[OF convex_scaling[OF assms], of a c] by auto qed -lemma pos_is_convex: "convex {0 :: real <..}" - unfolding convex_alt -proof safe - fix y x \ :: real - assume *: "y > 0" "x > 0" "\ \ 0" "\ \ 1" - { - assume "\ = 0" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" - by simp - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by simp - } - moreover - { - assume "\ = 1" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by simp - } - moreover - { - assume "\ \ 1" "\ \ 0" - then have "\ > 0" "(1 - \) > 0" - using * by auto - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by (auto simp: add_pos_pos) - } - ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" - by fastforce -qed - lemma convex_on_sum: fixes a :: "'a \ real" and y :: "'a \ 'b::real_vector" and f :: "'b \ real" assumes "finite s" "s \ {}" and "convex_on C f" and "convex C" and "(\ i \ s. a i) = 1" and "\i. i \ s \ a i \ 0" and "\i. i \ s \ y i \ C" shows "f (\ i \ s. a i *\<^sub>R y i) \ (\ i \ s. a i * f (y i))" using assms proof (induct s arbitrary: a rule: finite_ne_induct) case (singleton i) then have ai: "a i = 1" by auto then show ?case by auto next case (insert i s) then have "convex_on C f" by simp from this[unfolded convex_on_def, rule_format] have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by simp show ?case proof (cases "a i = 1") case True then have "(\ j \ s. a j) = 0" using insert by auto then have "\j. j \ s \ a j = 0" using insert by (fastforce simp: sum_nonneg_eq_0_iff) then show ?thesis using insert by auto next case False from insert have yai: "y i \ C" "a i \ 0" by auto have fis: "finite (insert i s)" using insert by auto then have ai1: "a i \ 1" using sum_nonneg_leq_bound[of "insert i s" a] insert by simp then have "a i < 1" using False by auto then have i0: "1 - a i > 0" by auto let ?a = "\j. a j / (1 - a i)" have a_nonneg: "?a j \ 0" if "j \ s" for j using i0 insert that by fastforce have "(\ j \ insert i s. a j) = 1" using insert by auto then have "(\ j \ s. a j) = 1 - a i" using sum.insert insert by fastforce then have "(\ j \ s. a j) / (1 - a i) = 1" using i0 by auto then have a1: "(\ j \ s. ?a j) = 1" unfolding sum_divide_distrib by simp have "convex C" using insert by auto then have asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" using insert convex_sum [OF \finite s\ \convex C\ a1 a_nonneg] by auto have asum_le: "f (\ j \ s. ?a j *\<^sub>R y j) \ (\ j \ s. ?a j * f (y j))" using a_nonneg a1 insert by blast have "f (\ j \ insert i s. a j *\<^sub>R y j) = f ((\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using sum.insert[of s i "\ j. a j *\<^sub>R y j", OF \finite s\ \i \ s\] insert by (auto simp only: add.commute) also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using i0 by auto also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" using scaleR_right.sum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" s, symmetric] by (auto simp: algebra_simps) also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" by (auto simp: divide_inverse) also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ s. ?a j *\<^sub>R y j)) + a i * f (y i)" using conv[of "y i" "(\ j \ s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] by (auto simp: add.commute) also have "\ \ (1 - a i) * (\ j \ s. ?a j * f (y j)) + a i * f (y i)" using add_right_mono [OF mult_left_mono [of _ _ "1 - a i", OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"] by simp also have "\ = (\ j \ s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)" unfolding sum_distrib_left[of "1 - a i" "\ j. ?a j * f (y j)"] using i0 by auto also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" using i0 by auto also have "\ = (\ j \ insert i s. a j * f (y j))" using insert by auto finally show ?thesis by simp qed qed lemma convex_on_alt: fixes C :: "'a::real_vector set" assumes "convex C" shows "convex_on C f \ (\x \ C. \ y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" proof safe fix x y fix \ :: real assume *: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" from this[unfolded convex_on_def, rule_format] have "0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" for u v by auto from this [of "\" "1 - \", simplified] * show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by auto next assume *: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" { fix x y fix u v :: real assume **: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" then have[simp]: "1 - u = v" by auto from *[rule_format, of x y u] have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" using ** by auto } then show "convex_on C f" unfolding convex_on_def by auto qed lemma convex_on_diff: fixes f :: "real \ real" assumes f: "convex_on I f" and I: "x \ I" "y \ I" and t: "x < t" "t < y" shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" and "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" proof - define a where "a \ (t - y) / (x - y)" with t have "0 \ a" "0 \ 1 - a" by (auto simp: field_simps) with f \x \ I\ \y \ I\ have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" by (auto simp: convex_on_def) have "a * x + (1 - a) * y = a * (x - y) + y" by (simp add: field_simps) also have "\ = t" unfolding a_def using \x < t\ \t < y\ by simp finally have "f t \ a * f x + (1 - a) * f y" using cvx by simp also have "\ = a * (f x - f y) + f y" by (simp add: field_simps) finally have "f t - f y \ a * (f x - f y)" by simp with t show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" by (simp add: le_divide_eq divide_le_eq field_simps a_def) with t show "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" by (simp add: le_divide_eq divide_le_eq field_simps) qed lemma pos_convex_function: fixes f :: "real \ real" assumes "convex C" and leq: "\x y. x \ C \ y \ C \ f' x * (y - x) \ f y - f x" shows "convex_on C f" unfolding convex_on_alt[OF assms(1)] using assms proof safe fix x y \ :: real let ?x = "\ *\<^sub>R x + (1 - \) *\<^sub>R y" assume *: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" then have "1 - \ \ 0" by auto then have xpos: "?x \ C" using * unfolding convex_alt by fastforce have geq: "\ * (f x - f ?x) + (1 - \) * (f y - f ?x) \ \ * f' ?x * (x - ?x) + (1 - \) * f' ?x * (y - ?x)" using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \\ \ 0\] mult_left_mono [OF leq [OF xpos *(3)] \1 - \ \ 0\]] by auto then have "\ * f x + (1 - \) * f y - f ?x \ 0" by (auto simp: field_simps) then show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" using convex_on_alt by auto qed lemma atMostAtLeast_subset_convex: fixes C :: "real set" assumes "convex C" and "x \ C" "y \ C" "x < y" shows "{x .. y} \ C" proof safe fix z assume z: "z \ {x .. y}" have less: "z \ C" if *: "x < z" "z < y" proof - let ?\ = "(y - z) / (y - x)" have "0 \ ?\" "?\ \ 1" using assms * by (auto simp: field_simps) then have comb: "?\ * x + (1 - ?\) * y \ C" using assms iffD1[OF convex_alt, rule_format, of C y x ?\] by (simp add: algebra_simps) have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" by (auto simp: field_simps) also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" using assms by (simp only: add_divide_distrib) (auto simp: field_simps) also have "\ = z" using assms by (auto simp: field_simps) finally show ?thesis using comb by auto qed show "z \ C" using z less assms by (auto simp: le_less) qed lemma f''_imp_f': fixes f :: "real \ real" assumes "convex C" and f': "\x. x \ C \ DERIV f x :> (f' x)" and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" and pos: "\x. x \ C \ f'' x \ 0" and x: "x \ C" and y: "y \ C" shows "f' x * (y - x) \ f y - f x" using assms proof - have less_imp: "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" if *: "x \ C" "y \ C" "y > x" for x y :: real proof - from * have ge: "y - x > 0" "y - x \ 0" by auto from * have le: "x - y < 0" "x - y \ 0" by auto then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \y \ C\ \x < y\], THEN f', THEN MVT2[OF \x < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] by auto then have "z1 \ C" using atMostAtLeast_subset_convex \convex C\ \x \ C\ \y \ C\ \x < y\ by fastforce from z1 have z1': "f x - f y = (x - y) * f' z1" by (simp add: field_simps) obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \z1 \ C\ \y \ C\ \z1 < y\], THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto have "f' y - (f x - f y) / (x - y) = f' y - f' z1" using * z1' by auto also have "\ = (y - z1) * f'' z3" using z3 by auto finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" by simp have A': "y - z1 \ 0" using z1 by auto have "z3 \ C" using z3 * atMostAtLeast_subset_convex \convex C\ \x \ C\ \z1 \ C\ \x < z1\ by fastforce then have B': "f'' z3 \ 0" using assms by auto from A' B' have "(y - z1) * f'' z3 \ 0" by auto from cool' this have "f' y - (f x - f y) / (x - y) \ 0" by auto from mult_right_mono_neg[OF this le(2)] have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" by (simp add: algebra_simps) then have "f' y * (x - y) - (f x - f y) \ 0" using le by auto then have res: "f' y * (x - y) \ f x - f y" by auto have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" using * z1 by auto also have "\ = (z1 - x) * f'' z2" using z2 by auto finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" by simp have A: "z1 - x \ 0" using z1 by auto have "z2 \ C" using z2 z1 * atMostAtLeast_subset_convex \convex C\ \z1 \ C\ \y \ C\ \z1 < y\ by fastforce then have B: "f'' z2 \ 0" using assms by auto from A B have "(z1 - x) * f'' z2 \ 0" by auto with cool have "(f y - f x) / (y - x) - f' x \ 0" by auto from mult_right_mono[OF this ge(2)] have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" by (simp add: algebra_simps) then have "f y - f x - f' x * (y - x) \ 0" using ge by auto then show "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" using res by auto qed show ?thesis proof (cases "x = y") case True with x y show ?thesis by auto next case False with less_imp x y show ?thesis by (auto simp: neq_iff) qed qed lemma f''_ge0_imp_convex: fixes f :: "real \ real" assumes conv: "convex C" and f': "\x. x \ C \ DERIV f x :> (f' x)" and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" and pos: "\x. x \ C \ f'' x \ 0" shows "convex_on C f" using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function by fastforce lemma minus_log_convex: fixes b :: real assumes "b > 1" shows "convex_on {0 <..} (\ x. - log b x)" proof - have "\z. z > 0 \ DERIV (log b) z :> 1 / (ln b * z)" using DERIV_log by auto then have f': "\z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" by (auto simp: DERIV_minus) have "\z::real. z > 0 \ DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto from this[THEN DERIV_cmult, of _ "- 1 / ln b"] have "\z::real. z > 0 \ DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" by auto then have f''0: "\z::real. z > 0 \ DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" unfolding inverse_eq_divide by (auto simp: mult.assoc) have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" using \b > 1\ by (auto intro!: less_imp_le) - from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] + from f''_ge0_imp_convex[OF convex_real_interval(3), unfolded greaterThan_iff, OF f' f''0 f''_ge0] show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \Convexity of real functions\ lemma convex_on_realI: assumes "connected A" and "\x. x \ A \ (f has_real_derivative f' x) (at x)" and "\x y. x \ A \ y \ A \ x \ y \ f' x \ f' y" shows "convex_on A f" proof (rule convex_on_linorderI) fix t x y :: real assume t: "t > 0" "t < 1" assume xy: "x \ A" "y \ A" "x < y" define z where "z = (1 - t) * x + t * y" with \connected A\ and xy have ivl: "{x..y} \ A" using connected_contains_Icc by blast from xy t have xz: "z > x" by (simp add: z_def algebra_simps) have "y - z = (1 - t) * (y - x)" by (simp add: z_def algebra_simps) also from xy t have "\ > 0" by (intro mult_pos_pos) simp_all finally have yz: "z < y" by simp from assms xz yz ivl t have "\\. \ > x \ \ < z \ f z - f x = (z - x) * f' \" by (intro MVT2) (auto intro!: assms(2)) then obtain \ where \: "\ > x" "\ < z" "f' \ = (f z - f x) / (z - x)" by auto from assms xz yz ivl t have "\\. \ > z \ \ < y \ f y - f z = (y - z) * f' \" by (intro MVT2) (auto intro!: assms(2)) then obtain \ where \: "\ > z" "\ < y" "f' \ = (f y - f z) / (y - z)" by auto from \(3) have "(f y - f z) / (y - z) = f' \" .. also from \ \ ivl have "\ \ A" "\ \ A" by auto with \ \ have "f' \ \ f' \" by (intro assms(3)) auto also from \(3) have "f' \ = (f z - f x) / (z - x)" . finally have "(f y - f z) * (z - x) \ (f z - f x) * (y - z)" using xz yz by (simp add: field_simps) also have "z - x = t * (y - x)" by (simp add: z_def algebra_simps) also have "y - z = (1 - t) * (y - x)" by (simp add: z_def algebra_simps) finally have "(f y - f z) * t \ (f z - f x) * (1 - t)" using xy by simp then show "(1 - t) * f x + t * f y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" by (simp add: z_def algebra_simps) qed lemma convex_on_inverse: assumes "A \ {0<..}" shows "convex_on A (inverse :: real \ real)" proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\x. -inverse (x^2)"]) fix u v :: real assume "u \ {0<..}" "v \ {0<..}" "u \ v" with assms show "-inverse (u^2) \ -inverse (v^2)" by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) qed (insert assms, auto intro!: derivative_eq_intros simp: field_split_simps power2_eq_square) lemma convex_onD_Icc': assumes "convex_on {x..y} f" "c \ {x..y}" defines "d \ y - x" shows "f c \ (f y - f x) / d * (c - x) + f x" proof (cases x y rule: linorder_cases) case less then have d: "d > 0" by (simp add: d_def) from assms(2) less have A: "0 \ (c - x) / d" "(c - x) / d \ 1" by (simp_all add: d_def field_split_simps) have "f c = f (x + (c - x) * 1)" by simp also from less have "1 = ((y - x) / d)" by (simp add: d_def) also from d have "x + (c - x) * \ = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" by (simp add: field_simps) also have "f \ \ (1 - (c - x) / d) * f x + (c - x) / d * f y" using assms less by (intro convex_onD_Icc) simp_all also from d have "\ = (f y - f x) / d * (c - x) + f x" by (simp add: field_simps) finally show ?thesis . qed (insert assms(2), simp_all) lemma convex_onD_Icc'': assumes "convex_on {x..y} f" "c \ {x..y}" defines "d \ y - x" shows "f c \ (f x - f y) / d * (y - c) + f y" proof (cases x y rule: linorder_cases) case less then have d: "d > 0" by (simp add: d_def) from assms(2) less have A: "0 \ (y - c) / d" "(y - c) / d \ 1" by (simp_all add: d_def field_split_simps) have "f c = f (y - (y - c) * 1)" by simp also from less have "1 = ((y - x) / d)" by (simp add: d_def) also from d have "y - (y - c) * \ = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" by (simp add: field_simps) also have "f \ \ (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) also from d have "\ = (f x - f y) / d * (y - c) + f y" by (simp add: field_simps) finally show ?thesis . qed (insert assms(2), simp_all) lemma convex_translation_eq [simp]: "convex ((+) a ` s) \ convex s" by (metis convex_translation translation_galois) lemma convex_translation_subtract_eq [simp]: "convex ((\b. b - a) ` s) \ convex s" using convex_translation_eq [of "- a"] by (simp cong: image_cong_simp) lemma convex_linear_image_eq [simp]: fixes f :: "'a::real_vector \ 'b::real_vector" shows "\linear f; inj f\ \ convex (f ` s) \ convex s" by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq) -lemma fst_linear: "linear fst" - unfolding linear_iff by (simp add: algebra_simps) - -lemma snd_linear: "linear snd" - unfolding linear_iff by (simp add: algebra_simps) - lemma fst_snd_linear: "linear (\(x,y). x + y)" unfolding linear_iff by (simp add: algebra_simps) lemma vector_choose_size: assumes "0 \ c" obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c" proof - obtain a::'a where "a \ 0" using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce then show ?thesis by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms) qed lemma vector_choose_dist: assumes "0 \ c" obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c" by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size) lemma sum_delta'': fixes s::"'a::real_vector set" assumes "finite s" shows "(\x\s. (if y = x then f x else 0) *\<^sub>R x) = (if y\s then (f y) *\<^sub>R y else 0)" proof - have *: "\x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)" by auto show ?thesis unfolding * using sum.delta[OF assms, of y "\x. f x *\<^sub>R x"] by auto qed lemma dist_triangle_eq: fixes x y z :: "'a::real_inner" shows "dist x z = dist x y + dist y z \ norm (x - y) *\<^sub>R (y - z) = norm (y - z) *\<^sub>R (x - y)" proof - have *: "x - y + (y - z) = x - z" by auto show ?thesis unfolding dist_norm norm_triangle_eq[of "x - y" "y - z", unfolded *] by (auto simp:norm_minus_commute) qed subsection \Cones\ definition\<^marker>\tag important\ cone :: "'a::real_vector set \ bool" where "cone s \ (\x\s. \c\0. c *\<^sub>R x \ s)" lemma cone_empty[intro, simp]: "cone {}" unfolding cone_def by auto lemma cone_univ[intro, simp]: "cone UNIV" unfolding cone_def by auto lemma cone_Inter[intro]: "\s\f. cone s \ cone (\f)" unfolding cone_def by auto lemma subspace_imp_cone: "subspace S \ cone S" by (simp add: cone_def subspace_scale) subsubsection \Conic hull\ lemma cone_cone_hull: "cone (cone hull s)" unfolding hull_def by auto lemma cone_hull_eq: "cone hull s = s \ cone s" apply (rule hull_eq) using cone_Inter unfolding subset_eq apply auto done lemma mem_cone: assumes "cone S" "x \ S" "c \ 0" shows "c *\<^sub>R x \ S" using assms cone_def[of S] by auto lemma cone_contains_0: assumes "cone S" shows "S \ {} \ 0 \ S" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ S" using assms mem_cone[of S a 0] by auto } then show ?thesis by auto qed lemma cone_0: "cone {0}" unfolding cone_def by auto lemma cone_Union[intro]: "(\s\f. cone s) \ cone (\f)" unfolding cone_def by blast lemma cone_iff: assumes "S \ {}" shows "cone S \ 0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" proof - { assume "cone S" { fix c :: real assume "c > 0" { fix x assume "x \ S" then have "x \ ((*\<^sub>R) c) ` S" unfolding image_def using \cone S\ \c>0\ mem_cone[of S x "1/c"] exI[of "(\t. t \ S \ x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] by auto } moreover { fix x assume "x \ ((*\<^sub>R) c) ` S" then have "x \ S" using \cone S\ \c > 0\ unfolding cone_def image_def \c > 0\ by auto } ultimately have "((*\<^sub>R) c) ` S = S" by blast } then have "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" using \cone S\ cone_contains_0[of S] assms by auto } moreover { assume a: "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" { fix x assume "x \ S" fix c1 :: real assume "c1 \ 0" then have "c1 = 0 \ c1 > 0" by auto then have "c1 *\<^sub>R x \ S" using a \x \ S\ by auto } then have "cone S" unfolding cone_def by auto } ultimately show ?thesis by blast qed lemma cone_hull_empty: "cone hull {} = {}" by (metis cone_empty cone_hull_eq) lemma cone_hull_empty_iff: "S = {} \ cone hull S = {}" by (metis bot_least cone_hull_empty hull_subset xtrans(5)) lemma cone_hull_contains_0: "S \ {} \ 0 \ cone hull S" using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S] by auto lemma mem_cone_hull: assumes "x \ S" "c \ 0" shows "c *\<^sub>R x \ cone hull S" by (metis assms cone_cone_hull hull_inc mem_cone) proposition cone_hull_expl: "cone hull S = {c *\<^sub>R x | c x. c \ 0 \ x \ S}" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?rhs" then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" by auto fix c :: real assume c: "c \ 0" then have "c *\<^sub>R x = (c * cx) *\<^sub>R xx" using x by (simp add: algebra_simps) moreover have "c * cx \ 0" using c x by auto ultimately have "c *\<^sub>R x \ ?rhs" using x by auto } then have "cone ?rhs" unfolding cone_def by auto then have "?rhs \ Collect cone" unfolding mem_Collect_eq by auto { fix x assume "x \ S" then have "1 *\<^sub>R x \ ?rhs" apply auto apply (rule_tac x = 1 in exI, auto) done then have "x \ ?rhs" by auto } then have "S \ ?rhs" by auto then have "?lhs \ ?rhs" using \?rhs \ Collect cone\ hull_minimal[of S "?rhs" "cone"] by auto moreover { fix x assume "x \ ?rhs" then obtain cx :: real and xx where x: "x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" by auto then have "xx \ cone hull S" using hull_subset[of S] by auto then have "x \ ?lhs" using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto } ultimately show ?thesis by auto qed lemma convex_cone: "convex s \ cone s \ (\x\s. \y\s. (x + y) \ s) \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" (is "?lhs = ?rhs") proof - { fix x y assume "x\s" "y\s" and ?lhs then have "2 *\<^sub>R x \s" "2 *\<^sub>R y \ s" unfolding cone_def by auto then have "x + y \ s" using \?lhs\[unfolded convex_def, THEN conjunct1] apply (erule_tac x="2*\<^sub>R x" in ballE) apply (erule_tac x="2*\<^sub>R y" in ballE) apply (erule_tac x="1/2" in allE, simp) apply (erule_tac x="1/2" in allE, auto) done } then show ?thesis unfolding convex_def cone_def by blast qed subsection\<^marker>\tag unimportant\ \Connectedness of convex sets\ -lemma connectedD: - "connected S \ open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S = {} \ B \ S = {}" - by (rule Topological_Spaces.topological_space_class.connectedD) - lemma convex_connected: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "connected S" proof (rule connectedI) fix A B assume "open A" "open B" "A \ B \ S = {}" "S \ A \ B" moreover assume "A \ S \ {}" "B \ S \ {}" then obtain a b where a: "a \ A" "a \ S" and b: "b \ B" "b \ S" by auto define f where [abs_def]: "f u = u *\<^sub>R a + (1 - u) *\<^sub>R b" for u then have "continuous_on {0 .. 1} f" by (auto intro!: continuous_intros) then have "connected (f ` {0 .. 1})" by (auto intro!: connected_continuous_image) note connectedD[OF this, of A B] moreover have "a \ A \ f ` {0 .. 1}" using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def) moreover have "b \ B \ f ` {0 .. 1}" using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def) moreover have "f ` {0 .. 1} \ S" using \convex S\ a b unfolding convex_def f_def by auto ultimately show False by auto qed corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_connected) lemma convex_prod: assumes "\i. i \ Basis \ convex {x. P i x}" shows "convex {x. \i\Basis. P i (x\i)}" using assms unfolding convex_def by (auto simp: inner_add_left) lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\i\Basis. 0 \ x\i)}" by (rule convex_prod) (simp flip: atLeast_def) subsection \Convex hull\ lemma convex_convex_hull [iff]: "convex (convex hull s)" unfolding hull_def using convex_Inter[of "{t. convex t \ s \ t}"] by auto lemma convex_hull_subset: "s \ convex hull t \ convex hull s \ convex hull t" by (simp add: subset_hull) lemma convex_hull_eq: "convex hull s = s \ convex s" by (metis convex_convex_hull hull_same) subsubsection\<^marker>\tag unimportant\ \Convex hull is "preserved" by a linear function\ lemma convex_hull_linear_image: assumes f: "linear f" shows "f ` (convex hull s) = convex hull (f ` s)" proof show "convex hull (f ` s) \ f ` (convex hull s)" by (intro hull_minimal image_mono hull_subset convex_linear_image assms convex_convex_hull) show "f ` (convex hull s) \ convex hull (f ` s)" proof (unfold image_subset_iff_subset_vimage, rule hull_minimal) show "s \ f -` (convex hull (f ` s))" by (fast intro: hull_inc) show "convex (f -` (convex hull (f ` s)))" by (intro convex_linear_vimage [OF f] convex_convex_hull) qed qed lemma in_convex_hull_linear_image: assumes "linear f" and "x \ convex hull s" shows "f x \ convex hull (f ` s)" using convex_hull_linear_image[OF assms(1)] assms(2) by auto lemma convex_hull_Times: "convex hull (s \ t) = (convex hull s) \ (convex hull t)" proof show "convex hull (s \ t) \ (convex hull s) \ (convex hull t)" by (intro hull_minimal Sigma_mono hull_subset convex_Times convex_convex_hull) have "(x, y) \ convex hull (s \ t)" if x: "x \ convex hull s" and y: "y \ convex hull t" for x y proof (rule hull_induct [OF x], rule hull_induct [OF y]) fix x y assume "x \ s" and "y \ t" then show "(x, y) \ convex hull (s \ t)" by (simp add: hull_inc) next fix x let ?S = "((\y. (0, y)) -` (\p. (- x, 0) + p) ` (convex hull s \ t))" have "convex ?S" by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) also have "?S = {y. (x, y) \ convex hull (s \ t)}" by (auto simp: image_def Bex_def) finally show "convex {y. (x, y) \ convex hull (s \ t)}" . next show "convex {x. (x, y) \ convex hull s \ t}" proof - fix y let ?S = "((\x. (x, 0)) -` (\p. (0, - y) + p) ` (convex hull s \ t))" have "convex ?S" by (intro convex_linear_vimage convex_translation convex_convex_hull, simp add: linear_iff) also have "?S = {x. (x, y) \ convex hull (s \ t)}" by (auto simp: image_def Bex_def) finally show "convex {x. (x, y) \ convex hull (s \ t)}" . qed qed then show "(convex hull s) \ (convex hull t) \ convex hull (s \ t)" unfolding subset_eq split_paired_Ball_Sigma by blast qed subsubsection\<^marker>\tag unimportant\ \Stepping theorems for convex hulls of finite sets\ lemma convex_hull_empty[simp]: "convex hull {} = {}" by (rule hull_unique) auto lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" by (rule hull_unique) auto lemma convex_hull_insert: fixes S :: "'a::real_vector set" assumes "S \ {}" shows "convex hull (insert a S) = {x. \u\0. \v\0. \b. (u + v = 1) \ b \ (convex hull S) \ (x = u *\<^sub>R a + v *\<^sub>R b)}" (is "_ = ?hull") proof (intro equalityI hull_minimal subsetI) fix x assume "x \ insert a S" then have "\u\0. \v\0. u + v = 1 \ (\b. b \ convex hull S \ x = u *\<^sub>R a + v *\<^sub>R b)" unfolding insert_iff proof assume "x = a" then show ?thesis by (rule_tac x=1 in exI) (use assms hull_subset in fastforce) next assume "x \ S" with hull_subset[of S convex] show ?thesis by force qed then show "x \ ?hull" by simp next fix x assume "x \ ?hull" then obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "x = u *\<^sub>R a + v *\<^sub>R b" by auto have "a \ convex hull insert a S" "b \ convex hull insert a S" using hull_mono[of S "insert a S" convex] hull_mono[of "{a}" "insert a S" convex] and obt(4) by auto then show "x \ convex hull insert a S" unfolding obt(5) using obt(1-3) by (rule convexD [OF convex_convex_hull]) next show "convex ?hull" proof (rule convexI) fix x y u v assume as: "(0::real) \ u" "0 \ v" "u + v = 1" and x: "x \ ?hull" and y: "y \ ?hull" from x obtain u1 v1 b1 where obt1: "u1\0" "v1\0" "u1 + v1 = 1" "b1 \ convex hull S" and xeq: "x = u1 *\<^sub>R a + v1 *\<^sub>R b1" by auto from y obtain u2 v2 b2 where obt2: "u2\0" "v2\0" "u2 + v2 = 1" "b2 \ convex hull S" and yeq: "y = u2 *\<^sub>R a + v2 *\<^sub>R b2" by auto have *: "\(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp: algebra_simps) have "\b \ convex hull S. u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" proof (cases "u * v1 + v * v2 = 0") case True have *: "\(x::'a) s1 s2. x - s1 *\<^sub>R x - s2 *\<^sub>R x = ((1::real) - (s1 + s2)) *\<^sub>R x" by (auto simp: algebra_simps) have eq0: "u * v1 = 0" "v * v2 = 0" using True mult_nonneg_nonneg[OF \u\0\ \v1\0\] mult_nonneg_nonneg[OF \v\0\ \v2\0\] by arith+ then have "u * u1 + v * u2 = 1" using as(3) obt1(3) obt2(3) by auto then show ?thesis using "*" eq0 as obt1(4) xeq yeq by auto next case False have "1 - (u * u1 + v * u2) = (u + v) - (u * u1 + v * u2)" using as(3) obt1(3) obt2(3) by (auto simp: field_simps) also have "\ = u * (v1 + u1 - u1) + v * (v2 + u2 - u2)" using as(3) obt1(3) obt2(3) by (auto simp: field_simps) also have "\ = u * v1 + v * v2" by simp finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto let ?b = "((u * v1) / (u * v1 + v * v2)) *\<^sub>R b1 + ((v * v2) / (u * v1 + v * v2)) *\<^sub>R b2" have zeroes: "0 \ u * v1 + v * v2" "0 \ u * v1" "0 \ u * v1 + v * v2" "0 \ v * v2" using as(1,2) obt1(1,2) obt2(1,2) by auto show ?thesis proof show "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (?b - (u * u1) *\<^sub>R ?b - (v * u2) *\<^sub>R ?b)" unfolding xeq yeq * ** using False by (auto simp: scaleR_left_distrib scaleR_right_distrib) show "?b \ convex hull S" using False zeroes obt1(4) obt2(4) by (auto simp: convexD [OF convex_convex_hull] scaleR_left_distrib scaleR_right_distrib add_divide_distrib[symmetric] zero_le_divide_iff) qed qed then obtain b where b: "b \ convex hull S" "u *\<^sub>R x + v *\<^sub>R y = (u * u1) *\<^sub>R a + (v * u2) *\<^sub>R a + (b - (u * u1) *\<^sub>R b - (v * u2) *\<^sub>R b)" .. have u1: "u1 \ 1" unfolding obt1(3)[symmetric] and not_le using obt1(2) by auto have u2: "u2 \ 1" unfolding obt2(3)[symmetric] and not_le using obt2(2) by auto have "u1 * u + u2 * v \ max u1 u2 * u + max u1 u2 * v" proof (rule add_mono) show "u1 * u \ max u1 u2 * u" "u2 * v \ max u1 u2 * v" by (simp_all add: as mult_right_mono) qed also have "\ \ 1" unfolding distrib_left[symmetric] and as(3) using u1 u2 by auto finally have le1: "u1 * u + u2 * v \ 1" . show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" proof (intro CollectI exI conjI) show "0 \ u * u1 + v * u2" by (simp add: as(1) as(2) obt1(1) obt2(1)) show "0 \ 1 - u * u1 - v * u2" by (simp add: le1 diff_diff_add mult.commute) qed (use b in \auto simp: algebra_simps\) qed qed lemma convex_hull_insert_alt: "convex hull (insert a S) = (if S = {} then {a} else {(1 - u) *\<^sub>R a + u *\<^sub>R x |x u. 0 \ u \ u \ 1 \ x \ convex hull S})" apply (auto simp: convex_hull_insert) using diff_eq_eq apply fastforce by (metis add.group_left_neutral add_le_imp_le_diff diff_add_cancel) subsubsection\<^marker>\tag unimportant\ \Explicit expression for convex hull\ proposition convex_hull_indexed: fixes S :: "'a::real_vector set" shows "convex hull S = {y. \k u x. (\i\{1::nat .. k}. 0 \ u i \ x i \ S) \ (sum u {1..k} = 1) \ (\i = 1..k. u i *\<^sub>R x i) = y}" (is "?xyz = ?hull") proof (rule hull_unique [OF _ convexI]) show "S \ ?hull" by (clarsimp, rule_tac x=1 in exI, rule_tac x="\x. 1" in exI, auto) next fix T assume "S \ T" "convex T" then show "?hull \ T" by (blast intro: convex_sum) next fix x y u v assume uv: "0 \ u" "0 \ v" "u + v = (1::real)" assume xy: "x \ ?hull" "y \ ?hull" from xy obtain k1 u1 x1 where x [rule_format]: "\i\{1::nat..k1}. 0\u1 i \ x1 i \ S" "sum u1 {Suc 0..k1} = 1" "(\i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x" by auto from xy obtain k2 u2 x2 where y [rule_format]: "\i\{1::nat..k2}. 0\u2 i \ x2 i \ S" "sum u2 {Suc 0..k2} = 1" "(\i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y" by auto have *: "\P (x::'a) y s t i. (if P i then s else t) *\<^sub>R (if P i then x else y) = (if P i then s *\<^sub>R x else t *\<^sub>R y)" "{1..k1 + k2} \ {1..k1} = {1..k1}" "{1..k1 + k2} \ - {1..k1} = (\i. i + k1) ` {1..k2}" by auto have inj: "inj_on (\i. i + k1) {1..k2}" unfolding inj_on_def by auto let ?uu = "\i. if i \ {1..k1} then u * u1 i else v * u2 (i - k1)" let ?xx = "\i. if i \ {1..k1} then x1 i else x2 (i - k1)" show "u *\<^sub>R x + v *\<^sub>R y \ ?hull" proof (intro CollectI exI conjI ballI) show "0 \ ?uu i" "?xx i \ S" if "i \ {1..k1+k2}" for i using that by (auto simp add: le_diff_conv uv(1) x(1) uv(2) y(1)) show "(\i = 1..k1 + k2. ?uu i) = 1" "(\i = 1..k1 + k2. ?uu i *\<^sub>R ?xx i) = u *\<^sub>R x + v *\<^sub>R y" unfolding * sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] sum.reindex[OF inj] Collect_mem_eq o_def unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric] by (simp_all add: sum_distrib_left[symmetric] x(2,3) y(2,3) uv(3)) qed qed lemma convex_hull_finite: fixes S :: "'a::real_vector set" assumes "finite S" shows "convex hull S = {y. \u. (\x\S. 0 \ u x) \ sum u S = 1 \ sum (\x. u x *\<^sub>R x) S = y}" (is "?HULL = _") proof (rule hull_unique [OF _ convexI]; clarify) fix x assume "x \ S" then show "\u. (\x\S. 0 \ u x) \ sum u S = 1 \ (\x\S. u x *\<^sub>R x) = x" by (rule_tac x="\y. if x=y then 1 else 0" in exI) (auto simp: sum.delta'[OF assms] sum_delta''[OF assms]) next fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" fix ux assume ux [rule_format]: "\x\S. 0 \ ux x" "sum ux S = (1::real)" fix uy assume uy [rule_format]: "\x\S. 0 \ uy x" "sum uy S = (1::real)" have "0 \ u * ux x + v * uy x" if "x\S" for x by (simp add: that uv ux(1) uy(1)) moreover have "(\x\S. u * ux x + v * uy x) = 1" unfolding sum.distrib and sum_distrib_left[symmetric] ux(2) uy(2) using uv(3) by auto moreover have "(\x\S. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\x\S. ux x *\<^sub>R x) + v *\<^sub>R (\x\S. uy x *\<^sub>R x)" unfolding scaleR_left_distrib sum.distrib scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] by auto ultimately show "\uc. (\x\S. 0 \ uc x) \ sum uc S = 1 \ (\x\S. uc x *\<^sub>R x) = u *\<^sub>R (\x\S. ux x *\<^sub>R x) + v *\<^sub>R (\x\S. uy x *\<^sub>R x)" by (rule_tac x="\x. u * ux x + v * uy x" in exI, auto) qed (use assms in \auto simp: convex_explicit\) subsubsection\<^marker>\tag unimportant\ \Another formulation\ text "Formalized by Lars Schewe." lemma convex_hull_explicit: fixes p :: "'a::real_vector set" shows "convex hull p = {y. \S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ sum (\v. u v *\<^sub>R v) S = y}" (is "?lhs = ?rhs") proof - { fix x assume "x\?lhs" then obtain k u y where obt: "\i\{1::nat..k}. 0 \ u i \ y i \ p" "sum u {1..k} = 1" "(\i = 1..k. u i *\<^sub>R y i) = x" unfolding convex_hull_indexed by auto have fin: "finite {1..k}" by auto have fin': "\v. finite {i \ {1..k}. y i = v}" by auto { fix j assume "j\{1..k}" then have "y j \ p" "0 \ sum u {i. Suc 0 \ i \ i \ k \ y i = y j}" using obt(1)[THEN bspec[where x=j]] and obt(2) apply simp apply (rule sum_nonneg) using obt(1) apply auto done } moreover have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v}) = 1" unfolding sum.image_gen[OF fin, symmetric] using obt(2) by auto moreover have "(\v\y ` {1..k}. sum u {i \ {1..k}. y i = v} *\<^sub>R v) = x" using sum.image_gen[OF fin, of "\i. u i *\<^sub>R y i" y, symmetric] unfolding scaleR_left.sum using obt(3) by auto ultimately have "\S u. finite S \ S \ p \ (\x\S. 0 \ u x) \ sum u S = 1 \ (\v\S. u v *\<^sub>R v) = x" apply (rule_tac x="y ` {1..k}" in exI) apply (rule_tac x="\v. sum u {i\{1..k}. y i = v}" in exI, auto) done then have "x\?rhs" by auto } moreover { fix y assume "y\?rhs" then obtain S u where obt: "finite S" "S \ p" "\x\S. 0 \ u x" "sum u S = 1" "(\v\S. u v *\<^sub>R v) = y" by auto obtain f where f: "inj_on f {1..card S}" "f ` {1..card S} = S" using ex_bij_betw_nat_finite_1[OF obt(1)] unfolding bij_betw_def by auto { fix i :: nat assume "i\{1..card S}" then have "f i \ S" using f(2) by blast then have "0 \ u (f i)" "f i \ p" using obt(2,3) by auto } moreover have *: "finite {1..card S}" by auto { fix y assume "y\S" then obtain i where "i\{1..card S}" "f i = y" using f using image_iff[of y f "{1..card S}"] by auto then have "{x. Suc 0 \ x \ x \ card S \ f x = y} = {i}" apply auto using f(1)[unfolded inj_on_def] by (metis One_nat_def atLeastAtMost_iff) then have "card {x. Suc 0 \ x \ x \ card S \ f x = y} = 1" by auto then have "(\x\{x \ {1..card S}. f x = y}. u (f x)) = u y" "(\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y" by (auto simp: sum_constant_scaleR) } then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" unfolding sum.image_gen[OF *(1), of "\x. u (f x) *\<^sub>R f x" f] and sum.image_gen[OF *(1), of "\x. u (f x)" f] unfolding f using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x) *\<^sub>R f x)" "\v. u v *\<^sub>R v"] using sum.cong [of S S "\y. (\x\{x \ {1..card S}. f x = y}. u (f x))" u] unfolding obt(4,5) by auto ultimately have "\k u x. (\i\{1..k}. 0 \ u i \ x i \ p) \ sum u {1..k} = 1 \ (\i::nat = 1..k. u i *\<^sub>R x i) = y" apply (rule_tac x="card S" in exI) apply (rule_tac x="u \ f" in exI) apply (rule_tac x=f in exI, fastforce) done then have "y \ ?lhs" unfolding convex_hull_indexed by auto } ultimately show ?thesis unfolding set_eq_iff by blast qed subsubsection\<^marker>\tag unimportant\ \A stepping theorem for that expansion\ lemma convex_hull_finite_step: fixes S :: "'a::real_vector set" assumes "finite S" shows "(\u. (\x\insert a S. 0 \ u x) \ sum u (insert a S) = w \ sum (\x. u x *\<^sub>R x) (insert a S) = y) \ (\v\0. \u. (\x\S. 0 \ u x) \ sum u S = w - v \ sum (\x. u x *\<^sub>R x) S = y - v *\<^sub>R a)" (is "?lhs = ?rhs") proof (rule, case_tac[!] "a\S") assume "a \ S" then have *: "insert a S = S" by auto assume ?lhs then show ?rhs unfolding * by (rule_tac x=0 in exI, auto) next assume ?lhs then obtain u where u: "\x\insert a S. 0 \ u x" "sum u (insert a S) = w" "(\x\insert a S. u x *\<^sub>R x) = y" by auto assume "a \ S" then show ?rhs apply (rule_tac x="u a" in exI) using u(1)[THEN bspec[where x=a]] apply simp apply (rule_tac x=u in exI) using u[unfolded sum_clauses(2)[OF assms]] and \a\S\ apply auto done next assume "a \ S" then have *: "insert a S = S" by auto have fin: "finite (insert a S)" using assms by auto assume ?rhs then obtain v u where uv: "v\0" "\x\S. 0 \ u x" "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto show ?lhs apply (rule_tac x = "\x. (if a = x then v else 0) + u x" in exI) unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin] unfolding sum_clauses(2)[OF assms] using uv and uv(2)[THEN bspec[where x=a]] and \a\S\ apply auto done next assume ?rhs then obtain v u where uv: "v\0" "\x\S. 0 \ u x" "sum u S = w - v" "(\x\S. u x *\<^sub>R x) = y - v *\<^sub>R a" by auto moreover assume "a \ S" moreover have "(\x\S. if a = x then v else u x) = sum u S" "(\x\S. (if a = x then v else u x) *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" using \a \ S\ by (auto simp: intro!: sum.cong) ultimately show ?lhs by (rule_tac x="\x. if a = x then v else u x" in exI) (auto simp: sum_clauses(2)[OF assms]) qed subsubsection\<^marker>\tag unimportant\ \Hence some special cases\ lemma convex_hull_2: "convex hull {a,b} = {u *\<^sub>R a + v *\<^sub>R b | u v. 0 \ u \ 0 \ v \ u + v = 1}" proof - have *: "\u. (\x\{a, b}. 0 \ u x) \ 0 \ u a \ 0 \ u b" by auto have **: "finite {b}" by auto show ?thesis apply (simp add: convex_hull_finite) unfolding convex_hull_finite_step[OF **, of a 1, unfolded * conj_assoc] apply auto apply (rule_tac x=v in exI) apply (rule_tac x="1 - v" in exI, simp) apply (rule_tac x=u in exI, simp) apply (rule_tac x="\x. v" in exI, simp) done qed lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u. 0 \ u \ u \ 1}" unfolding convex_hull_2 proof (rule Collect_cong) have *: "\x y ::real. x + y = 1 \ x = 1 - y" by auto fix x show "(\v u. x = v *\<^sub>R a + u *\<^sub>R b \ 0 \ v \ 0 \ u \ v + u = 1) \ (\u. x = a + u *\<^sub>R (b - a) \ 0 \ u \ u \ 1)" unfolding * apply auto apply (rule_tac[!] x=u in exI) apply (auto simp: algebra_simps) done qed lemma convex_hull_3: "convex hull {a,b,c} = { u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c | u v w. 0 \ u \ 0 \ v \ 0 \ w \ u + v + w = 1}" proof - have fin: "finite {a,b,c}" "finite {b,c}" "finite {c}" by auto have *: "\x y z ::real. x + y + z = 1 \ x = 1 - y - z" by (auto simp: field_simps) show ?thesis unfolding convex_hull_finite[OF fin(1)] and convex_hull_finite_step[OF fin(2)] and * unfolding convex_hull_finite_step[OF fin(3)] apply (rule Collect_cong, simp) apply auto apply (rule_tac x=va in exI) apply (rule_tac x="u c" in exI, simp) apply (rule_tac x="1 - v - w" in exI, simp) apply (rule_tac x=v in exI, simp) apply (rule_tac x="\x. w" in exI, simp) done qed lemma convex_hull_3_alt: "convex hull {a,b,c} = {a + u *\<^sub>R (b - a) + v *\<^sub>R (c - a) | u v. 0 \ u \ 0 \ v \ u + v \ 1}" proof - have *: "\x y z ::real. x + y + z = 1 \ x = 1 - y - z" by auto show ?thesis unfolding convex_hull_3 apply (auto simp: *) apply (rule_tac x=v in exI) apply (rule_tac x=w in exI) apply (simp add: algebra_simps) apply (rule_tac x=u in exI) apply (rule_tac x=v in exI) apply (simp add: algebra_simps) done qed subsection\<^marker>\tag unimportant\ \Relations among closure notions and corresponding hulls\ lemma affine_imp_convex: "affine s \ convex s" unfolding affine_def convex_def by auto lemma convex_affine_hull [simp]: "convex (affine hull S)" by (simp add: affine_imp_convex) lemma subspace_imp_convex: "subspace s \ convex s" using subspace_imp_affine affine_imp_convex by auto lemma convex_hull_subset_span: "(convex hull s) \ (span s)" by (metis hull_minimal span_superset subspace_imp_convex subspace_span) lemma convex_hull_subset_affine_hull: "(convex hull s) \ (affine hull s)" by (metis affine_affine_hull affine_imp_convex hull_minimal hull_subset) lemma aff_dim_convex_hull: fixes S :: "'n::euclidean_space set" shows "aff_dim (convex hull S) = aff_dim S" using aff_dim_affine_hull[of S] convex_hull_subset_affine_hull[of S] hull_subset[of S "convex"] aff_dim_subset[of S "convex hull S"] aff_dim_subset[of "convex hull S" "affine hull S"] by auto subsection \Caratheodory's theorem\ lemma convex_hull_caratheodory_aff_dim: fixes p :: "('a::euclidean_space) set" shows "convex hull p = {y. \s u. finite s \ s \ p \ card s \ aff_dim p + 1 \ (\x\s. 0 \ u x) \ sum u s = 1 \ sum (\v. u v *\<^sub>R v) s = y}" unfolding convex_hull_explicit set_eq_iff mem_Collect_eq proof (intro allI iffI) fix y let ?P = "\n. \s u. finite s \ card s = n \ s \ p \ (\x\s. 0 \ u x) \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" assume "\s u. finite s \ s \ p \ (\x\s. 0 \ u x) \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" then obtain N where "?P N" by auto then have "\n\N. (\k ?P k) \ ?P n" apply (rule_tac ex_least_nat_le, auto) done then obtain n where "?P n" and smallest: "\k ?P k" by blast then obtain s u where obt: "finite s" "card s = n" "s\p" "\x\s. 0 \ u x" "sum u s = 1" "(\v\s. u v *\<^sub>R v) = y" by auto have "card s \ aff_dim p + 1" proof (rule ccontr, simp only: not_le) assume "aff_dim p + 1 < card s" then have "affine_dependent s" using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3) by blast then obtain w v where wv: "sum w s = 0" "v\s" "w v \ 0" "(\v\s. w v *\<^sub>R v) = 0" using affine_dependent_explicit_finite[OF obt(1)] by auto define i where "i = (\v. (u v) / (- w v)) ` {v\s. w v < 0}" define t where "t = Min i" have "\x\s. w x < 0" proof (rule ccontr, simp add: not_less) assume as:"\x\s. 0 \ w x" then have "sum w (s - {v}) \ 0" apply (rule_tac sum_nonneg, auto) done then have "sum w s > 0" unfolding sum.remove[OF obt(1) \v\s\] using as[THEN bspec[where x=v]] \v\s\ \w v \ 0\ by auto then show False using wv(1) by auto qed then have "i \ {}" unfolding i_def by auto then have "t \ 0" using Min_ge_iff[of i 0 ] and obt(1) unfolding t_def i_def using obt(4)[unfolded le_less] by (auto simp: divide_le_0_iff) have t: "\v\s. u v + t * w v \ 0" proof fix v assume "v \ s" then have v: "0 \ u v" using obt(4)[THEN bspec[where x=v]] by auto show "0 \ u v + t * w v" proof (cases "w v < 0") case False thus ?thesis using v \t\0\ by auto next case True then have "t \ u v / (- w v)" using \v\s\ unfolding t_def i_def apply (rule_tac Min_le) using obt(1) apply auto done then show ?thesis unfolding real_0_le_add_iff using pos_le_divide_eq[OF True[unfolded neg_0_less_iff_less[symmetric]]] by auto qed qed obtain a where "a \ s" and "t = (\v. (u v) / (- w v)) a" and "w a < 0" using Min_in[OF _ \i\{}\] and obt(1) unfolding i_def t_def by auto then have a: "a \ s" "u a + t * w a = 0" by auto have *: "\f. sum f (s - {a}) = sum f s - ((f a)::'b::ab_group_add)" unfolding sum.remove[OF obt(1) \a\s\] by auto have "(\v\s. u v + t * w v) = 1" unfolding sum.distrib wv(1) sum_distrib_left[symmetric] obt(5) by auto moreover have "(\v\s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y" unfolding sum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4) using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp ultimately have "?P (n - 1)" apply (rule_tac x="(s - {a})" in exI) apply (rule_tac x="\v. u v + t * w v" in exI) using obt(1-3) and t and a apply (auto simp: * scaleR_left_distrib) done then show False using smallest[THEN spec[where x="n - 1"]] by auto qed then show "\s u. finite s \ s \ p \ card s \ aff_dim p + 1 \ (\x\s. 0 \ u x) \ sum u s = 1 \ (\v\s. u v *\<^sub>R v) = y" using obt by auto qed auto lemma caratheodory_aff_dim: fixes p :: "('a::euclidean_space) set" shows "convex hull p = {x. \s. finite s \ s \ p \ card s \ aff_dim p + 1 \ x \ convex hull s}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" apply (subst convex_hull_caratheodory_aff_dim, clarify) apply (rule_tac x=s in exI) apply (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) done next show "?rhs \ ?lhs" using hull_mono by blast qed lemma convex_hull_caratheodory: fixes p :: "('a::euclidean_space) set" shows "convex hull p = {y. \s u. finite s \ s \ p \ card s \ DIM('a) + 1 \ (\x\s. 0 \ u x) \ sum u s = 1 \ sum (\v. u v *\<^sub>R v) s = y}" (is "?lhs = ?rhs") proof (intro set_eqI iffI) fix x assume "x \ ?lhs" then show "x \ ?rhs" apply (simp only: convex_hull_caratheodory_aff_dim Set.mem_Collect_eq) apply (erule ex_forward)+ using aff_dim_le_DIM [of p] apply simp done next fix x assume "x \ ?rhs" then show "x \ ?lhs" by (auto simp: convex_hull_explicit) qed theorem caratheodory: "convex hull p = {x::'a::euclidean_space. \s. finite s \ s \ p \ card s \ DIM('a) + 1 \ x \ convex hull s}" proof safe fix x assume "x \ convex hull p" then obtain s u where "finite s" "s \ p" "card s \ DIM('a) + 1" "\x\s. 0 \ u x" "sum u s = 1" "(\v\s. u v *\<^sub>R v) = x" unfolding convex_hull_caratheodory by auto then show "\s. finite s \ s \ p \ card s \ DIM('a) + 1 \ x \ convex hull s" apply (rule_tac x=s in exI) using hull_subset[of s convex] using convex_convex_hull[simplified convex_explicit, of s, THEN spec[where x=s], THEN spec[where x=u]] apply auto done next fix x s assume "finite s" "s \ p" "card s \ DIM('a) + 1" "x \ convex hull s" then show "x \ convex hull p" using hull_mono[OF \s\p\] by auto qed subsection\<^marker>\tag unimportant\\Some Properties of subset of standard basis\ lemma affine_hull_substd_basis: assumes "d \ Basis" shows "affine hull (insert 0 d) = {x::'a::euclidean_space. \i\Basis. i \ d \ x\i = 0}" (is "affine hull (insert 0 ?A) = ?B") proof - have *: "\A. (+) (0::'a) ` A = A" "\A. (+) (- (0::'a)) ` A = A" by auto show ?thesis unfolding affine_hull_insert_span_gen span_substd_basis[OF assms,symmetric] * .. qed lemma affine_hull_convex_hull [simp]: "affine hull (convex hull S) = affine hull S" by (metis Int_absorb1 Int_absorb2 convex_hull_subset_affine_hull hull_hull hull_mono hull_subset) subsection\<^marker>\tag unimportant\ \Moving and scaling convex hulls\ lemma convex_hull_set_plus: "convex hull (S + T) = convex hull S + convex hull T" unfolding set_plus_image apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_right_distrib) apply (simp add: convex_hull_Times) done lemma translation_eq_singleton_plus: "(\x. a + x) ` T = {a} + T" unfolding set_plus_def by auto lemma convex_hull_translation: "convex hull ((\x. a + x) ` S) = (\x. a + x) ` (convex hull S)" unfolding translation_eq_singleton_plus by (simp only: convex_hull_set_plus convex_hull_singleton) lemma convex_hull_scaling: "convex hull ((\x. c *\<^sub>R x) ` S) = (\x. c *\<^sub>R x) ` (convex hull S)" using linear_scaleR by (rule convex_hull_linear_image [symmetric]) lemma convex_hull_affinity: "convex hull ((\x. a + c *\<^sub>R x) ` S) = (\x. a + c *\<^sub>R x) ` (convex hull S)" by(simp only: image_image[symmetric] convex_hull_scaling convex_hull_translation) subsection\<^marker>\tag unimportant\ \Convexity of cone hulls\ lemma convex_cone_hull: assumes "convex S" shows "convex (cone hull S)" proof (rule convexI) fix x y assume xy: "x \ cone hull S" "y \ cone hull S" then have "S \ {}" using cone_hull_empty_iff[of S] by auto fix u v :: real assume uv: "u \ 0" "v \ 0" "u + v = 1" then have *: "u *\<^sub>R x \ cone hull S" "v *\<^sub>R y \ cone hull S" using cone_cone_hull[of S] xy cone_def[of "cone hull S"] by auto from * obtain cx :: real and xx where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" using cone_hull_expl[of S] by auto from * obtain cy :: real and yy where y: "v *\<^sub>R y = cy *\<^sub>R yy" "cy \ 0" "yy \ S" using cone_hull_expl[of S] by auto { assume "cx + cy \ 0" then have "u *\<^sub>R x = 0" and "v *\<^sub>R y = 0" using x y by auto then have "u *\<^sub>R x + v *\<^sub>R y = 0" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" using cone_hull_contains_0[of S] \S \ {}\ by auto } moreover { assume "cx + cy > 0" then have "(cx / (cx + cy)) *\<^sub>R xx + (cy / (cx + cy)) *\<^sub>R yy \ S" using assms mem_convex_alt[of S xx yy cx cy] x y by auto then have "cx *\<^sub>R xx + cy *\<^sub>R yy \ cone hull S" using mem_cone_hull[of "(cx/(cx+cy)) *\<^sub>R xx + (cy/(cx+cy)) *\<^sub>R yy" S "cx+cy"] \cx+cy>0\ by (auto simp: scaleR_right_distrib) then have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" using x y by auto } moreover have "cx + cy \ 0 \ cx + cy > 0" by auto ultimately show "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" by blast qed lemma cone_convex_hull: assumes "cone S" shows "cone (convex hull S)" proof (cases "S = {}") case True then show ?thesis by auto next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto { fix c :: real assume "c > 0" then have "(*\<^sub>R) c ` (convex hull S) = convex hull ((*\<^sub>R) c ` S)" using convex_hull_scaling[of _ S] by auto also have "\ = convex hull S" using * \c > 0\ by auto finally have "(*\<^sub>R) c ` (convex hull S) = convex hull S" by auto } then have "0 \ convex hull S" "\c. c > 0 \ ((*\<^sub>R) c ` (convex hull S)) = (convex hull S)" using * hull_subset[of S convex] by auto then show ?thesis using \S \ {}\ cone_iff[of "convex hull S"] by auto qed subsection \Radon's theorem\ text "Formalized by Lars Schewe." lemma Radon_ex_lemma: assumes "finite c" "affine_dependent c" shows "\u. sum u c = 0 \ (\v\c. u v \ 0) \ sum (\v. u v *\<^sub>R v) c = 0" proof - from assms(2)[unfolded affine_dependent_explicit] obtain s u where "finite s" "s \ c" "sum u s = 0" "\v\s. u v \ 0" "(\v\s. u v *\<^sub>R v) = 0" by blast then show ?thesis apply (rule_tac x="\v. if v\s then u v else 0" in exI) unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric] apply (auto simp: Int_absorb1) done qed lemma Radon_s_lemma: assumes "finite s" and "sum f s = (0::real)" shows "sum f {x\s. 0 < f x} = - sum f {x\s. f x < 0}" proof - have *: "\x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x" by auto show ?thesis unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)] and sum.distrib[symmetric] and * using assms(2) by assumption qed lemma Radon_v_lemma: assumes "finite s" and "sum f s = 0" and "\x. g x = (0::real) \ f x = (0::'a::euclidean_space)" shows "(sum f {x\s. 0 < g x}) = - sum f {x\s. g x < 0}" proof - have *: "\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" using assms(3) by auto show ?thesis unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)] and sum.distrib[symmetric] and * using assms(2) apply assumption done qed lemma Radon_partition: assumes "finite c" "affine_dependent c" shows "\m p. m \ p = {} \ m \ p = c \ (convex hull m) \ (convex hull p) \ {}" proof - obtain u v where uv: "sum u c = 0" "v\c" "u v \ 0" "(\v\c. u v *\<^sub>R v) = 0" using Radon_ex_lemma[OF assms] by auto have fin: "finite {x \ c. 0 < u x}" "finite {x \ c. 0 > u x}" using assms(1) by auto define z where "z = inverse (sum u {x\c. u x > 0}) *\<^sub>R sum (\x. u x *\<^sub>R x) {x\c. u x > 0}" have "sum u {x \ c. 0 < u x} \ 0" proof (cases "u v \ 0") case False then have "u v < 0" by auto then show ?thesis proof (cases "\w\{x \ c. 0 < u x}. u w > 0") case True then show ?thesis using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto next case False then have "sum u c \ sum (\x. if x=v then u v else 0) c" apply (rule_tac sum_mono, auto) done then show ?thesis unfolding sum.delta[OF assms(1)] using uv(2) and \u v < 0\ and uv(1) by auto qed qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto) then have *: "sum u {x\c. u x > 0} > 0" unfolding less_le apply (rule_tac conjI) apply (rule_tac sum_nonneg, auto) done moreover have "sum u ({x \ c. 0 < u x} \ {x \ c. u x < 0}) = sum u c" "(\x\{x \ c. 0 < u x} \ {x \ c. u x < 0}. u x *\<^sub>R x) = (\x\c. u x *\<^sub>R x)" using assms(1) apply (rule_tac[!] sum.mono_neutral_left, auto) done then have "sum u {x \ c. 0 < u x} = - sum u {x \ c. 0 > u x}" "(\x\{x \ c. 0 < u x}. u x *\<^sub>R x) = - (\x\{x \ c. 0 > u x}. u x *\<^sub>R x)" unfolding eq_neg_iff_add_eq_0 using uv(1,4) by (auto simp: sum.union_inter_neutral[OF fin, symmetric]) moreover have "\x\{v \ c. u v < 0}. 0 \ inverse (sum u {x \ c. 0 < u x}) * - u x" apply rule apply (rule mult_nonneg_nonneg) using * apply auto done ultimately have "z \ convex hull {v \ c. u v \ 0}" unfolding convex_hull_explicit mem_Collect_eq apply (rule_tac x="{v \ c. u v < 0}" in exI) apply (rule_tac x="\y. inverse (sum u {x\c. u x > 0}) * - u y" in exI) using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def apply (auto simp: sum_negf sum_distrib_left[symmetric]) done moreover have "\x\{v \ c. 0 < u v}. 0 \ inverse (sum u {x \ c. 0 < u x}) * u x" apply rule apply (rule mult_nonneg_nonneg) using * apply auto done then have "z \ convex hull {v \ c. u v > 0}" unfolding convex_hull_explicit mem_Collect_eq apply (rule_tac x="{v \ c. 0 < u v}" in exI) apply (rule_tac x="\y. inverse (sum u {x\c. u x > 0}) * u y" in exI) using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def using * apply (auto simp: sum_negf sum_distrib_left[symmetric]) done ultimately show ?thesis apply (rule_tac x="{v\c. u v \ 0}" in exI) apply (rule_tac x="{v\c. u v > 0}" in exI, auto) done qed theorem Radon: assumes "affine_dependent c" obtains m p where "m \ c" "p \ c" "m \ p = {}" "(convex hull m) \ (convex hull p) \ {}" proof - from assms[unfolded affine_dependent_explicit] obtain s u where "finite s" "s \ c" "sum u s = 0" "\v\s. u v \ 0" "(\v\s. u v *\<^sub>R v) = 0" by blast then have *: "finite s" "affine_dependent s" and s: "s \ c" unfolding affine_dependent_explicit by auto from Radon_partition[OF *] obtain m p where "m \ p = {}" "m \ p = s" "convex hull m \ convex hull p \ {}" by blast then show ?thesis apply (rule_tac that[of p m]) using s apply auto done qed subsection \Helly's theorem\ lemma Helly_induct: fixes f :: "'a::euclidean_space set set" assumes "card f = n" and "n \ DIM('a) + 1" and "\s\f. convex s" "\t\f. card t = DIM('a) + 1 \ \t \ {}" shows "\f \ {}" using assms proof (induction n arbitrary: f) case 0 then show ?case by auto next case (Suc n) have "finite f" using \card f = Suc n\ by (auto intro: card_ge_0_finite) show "\f \ {}" proof (cases "n = DIM('a)") case True then show ?thesis by (simp add: Suc.prems(1) Suc.prems(4)) next case False have "\(f - {s}) \ {}" if "s \ f" for s proof (rule Suc.IH[rule_format]) show "card (f - {s}) = n" by (simp add: Suc.prems(1) \finite f\ that) show "DIM('a) + 1 \ n" using False Suc.prems(2) by linarith show "\t. \t \ f - {s}; card t = DIM('a) + 1\ \ \t \ {}" by (simp add: Suc.prems(4) subset_Diff_insert) qed (use Suc in auto) then have "\s\f. \x. x \ \(f - {s})" by blast then obtain X where X: "\s. s\f \ X s \ \(f - {s})" by metis show ?thesis proof (cases "inj_on X f") case False then obtain s t where "s\t" and st: "s\f" "t\f" "X s = X t" unfolding inj_on_def by auto then have *: "\f = \(f - {s}) \ \(f - {t})" by auto show ?thesis by (metis "*" X disjoint_iff_not_equal st) next case True then obtain m p where mp: "m \ p = {}" "m \ p = X ` f" "convex hull m \ convex hull p \ {}" using Radon_partition[of "X ` f"] and affine_dependent_biggerset[of "X ` f"] unfolding card_image[OF True] and \card f = Suc n\ using Suc(3) \finite f\ and False by auto have "m \ X ` f" "p \ X ` f" using mp(2) by auto then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" unfolding subset_image_iff by auto then have "f \ (g \ h) = f" by auto then have f: "f = g \ h" using inj_on_Un_image_eq_iff[of X f "g \ h"] and True unfolding mp(2)[unfolded image_Un[symmetric] gh] by auto have *: "g \ h = {}" using mp(1) unfolding gh using inj_on_image_Int[OF True gh(3,4)] by auto have "convex hull (X ` h) \ \g" "convex hull (X ` g) \ \h" by (rule hull_minimal; use X * f in \auto simp: Suc.prems(3) convex_Inter\)+ then show ?thesis unfolding f using mp(3)[unfolded gh] by blast qed qed qed theorem Helly: fixes f :: "'a::euclidean_space set set" assumes "card f \ DIM('a) + 1" "\s\f. convex s" and "\t. \t\f; card t = DIM('a) + 1\ \ \t \ {}" shows "\f \ {}" apply (rule Helly_induct) using assms apply auto done subsection \Epigraphs of convex functions\ definition\<^marker>\tag important\ "epigraph S (f :: _ \ real) = {xy. fst xy \ S \ f (fst xy) \ snd xy}" lemma mem_epigraph: "(x, y) \ epigraph S f \ x \ S \ f x \ y" unfolding epigraph_def by auto lemma convex_epigraph: "convex (epigraph S f) \ convex_on S f \ convex S" proof safe assume L: "convex (epigraph S f)" then show "convex_on S f" by (auto simp: convex_def convex_on_def epigraph_def) show "convex S" using L apply (clarsimp simp: convex_def convex_on_def epigraph_def) apply (erule_tac x=x in allE) apply (erule_tac x="f x" in allE, safe) apply (erule_tac x=y in allE) apply (erule_tac x="f y" in allE) apply (auto simp: ) done next assume "convex_on S f" "convex S" then show "convex (epigraph S f)" unfolding convex_def convex_on_def epigraph_def apply safe apply (rule_tac [2] y="u * f a + v * f aa" in order_trans) apply (auto intro!:mult_left_mono add_mono) done qed lemma convex_epigraphI: "convex_on S f \ convex S \ convex (epigraph S f)" unfolding convex_epigraph by auto lemma convex_epigraph_convex: "convex S \ convex_on S f \ convex(epigraph S f)" by (simp add: convex_epigraph) subsubsection\<^marker>\tag unimportant\ \Use this to derive general bound property of convex function\ lemma convex_on: assumes "convex S" shows "convex_on S f \ (\k u x. (\i\{1..k::nat}. 0 \ u i \ x i \ S) \ sum u {1..k} = 1 \ f (sum (\i. u i *\<^sub>R x i) {1..k}) \ sum (\i. u i * f(x i)) {1..k})" unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq unfolding fst_sum snd_sum fst_scaleR snd_scaleR apply safe apply (drule_tac x=k in spec) apply (drule_tac x=u in spec) apply (drule_tac x="\i. (x i, f (x i))" in spec) apply simp using assms[unfolded convex] apply simp apply (rule_tac y="\i = 1..k. u i * f (fst (x i))" in order_trans, force) apply (rule sum_mono) apply (erule_tac x=i in allE) unfolding real_scaleR_def apply (rule mult_left_mono) using assms[unfolded convex] apply auto done subsection\<^marker>\tag unimportant\ \A bound within a convex hull\ lemma convex_on_convex_hull_bound: assumes "convex_on (convex hull s) f" and "\x\s. f x \ b" shows "\x\ convex hull s. f x \ b" proof fix x assume "x \ convex hull s" then obtain k u v where obt: "\i\{1..k::nat}. 0 \ u i \ v i \ s" "sum u {1..k} = 1" "(\i = 1..k. u i *\<^sub>R v i) = x" unfolding convex_hull_indexed mem_Collect_eq by auto have "(\i = 1..k. u i * f (v i)) \ b" using sum_mono[of "{1..k}" "\i. u i * f (v i)" "\i. u i * b"] unfolding sum_distrib_right[symmetric] obt(2) mult_1 apply (drule_tac meta_mp) apply (rule mult_left_mono) using assms(2) obt(1) apply auto done then show "f x \ b" using assms(1)[unfolded convex_on[OF convex_convex_hull], rule_format, of k u v] unfolding obt(2-3) using obt(1) and hull_subset[unfolded subset_eq, rule_format, of _ s] by auto qed lemma inner_sum_Basis[simp]: "i \ Basis \ (\Basis) \ i = 1" by (simp add: inner_sum_left sum.If_cases inner_Basis) lemma convex_set_plus: assumes "convex S" and "convex T" shows "convex (S + T)" proof - have "convex (\x\ S. \y \ T. {x + y})" using assms by (rule convex_sums) moreover have "(\x\ S. \y \ T. {x + y}) = S + T" unfolding set_plus_def by auto finally show "convex (S + T)" . qed lemma convex_set_sum: assumes "\i. i \ A \ convex (B i)" shows "convex (\i\A. B i)" proof (cases "finite A") case True then show ?thesis using assms by induct (auto simp: convex_set_plus) qed auto lemma finite_set_sum: assumes "finite A" and "\i\A. finite (B i)" shows "finite (\i\A. B i)" using assms by (induct set: finite, simp, simp add: finite_set_plus) lemma box_eq_set_sum_Basis: shows "{x. \i\Basis. x\i \ B i} = (\i\Basis. image (\x. x *\<^sub>R i) (B i))" apply (subst set_sum_alt [OF finite_Basis], safe) apply (fast intro: euclidean_representation [symmetric]) apply (subst inner_sum_left) apply (rename_tac f) apply (subgoal_tac "(\x\Basis. f x \ i) = f i \ i") apply (drule (1) bspec) apply clarsimp apply (frule sum.remove [OF finite_Basis]) apply (erule trans, simp) apply (rule sum.neutral, clarsimp) apply (frule_tac x=i in bspec, assumption) apply (drule_tac x=x in bspec, assumption, clarsimp) apply (cut_tac u=x and v=i in inner_Basis, assumption+) apply (rule ccontr, simp) done lemma convex_hull_set_sum: "convex hull (\i\A. B i) = (\i\A. convex hull (B i))" proof (cases "finite A") assume "finite A" then show ?thesis by (induct set: finite, simp, simp add: convex_hull_set_plus) qed simp end \ No newline at end of file diff --git a/src/HOL/Analysis/Convex_Euclidean_Space.thy b/src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy @@ -1,2280 +1,2274 @@ (* Title: HOL/Analysis/Convex_Euclidean_Space.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Convex Sets and Functions on (Normed) Euclidean Spaces\ theory Convex_Euclidean_Space imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets and Functions\ lemma aff_dim_cball: fixes a :: "'n::euclidean_space" assumes "e > 0" shows "aff_dim (cball a e) = int (DIM('n))" proof - have "(\x. a + x) ` (cball 0 e) \ cball a e" unfolding cball_def dist_norm by auto then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \ aff_dim (cball a e)" using aff_dim_translation_eq[of a "cball 0 e"] aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"] by auto moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) ultimately show ?thesis using aff_dim_le_DIM[of "cball a e"] by auto qed lemma aff_dim_open: fixes S :: "'n::euclidean_space set" assumes "open S" and "S \ {}" shows "aff_dim S = int (DIM('n))" proof - obtain x where "x \ S" using assms by auto then obtain e where e: "e > 0" "cball x e \ S" using open_contains_cball[of S] assms by auto then have "aff_dim (cball x e) \ aff_dim S" using aff_dim_subset by auto with e show ?thesis using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto qed lemma low_dim_interior: fixes S :: "'n::euclidean_space set" assumes "\ aff_dim S = int (DIM('n))" shows "interior S = {}" proof - have "aff_dim(interior S) \ aff_dim S" using interior_subset aff_dim_subset[of "interior S" S] by auto then show ?thesis using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto qed corollary empty_interior_lowdim: fixes S :: "'n::euclidean_space set" shows "dim S < DIM ('n) \ interior S = {}" by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) corollary aff_dim_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "interior S \ {} \ aff_dim S = DIM('a)" by (metis low_dim_interior) subsection \Relative interior of a set\ definition\<^marker>\tag important\ "rel_interior S = {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: "\S \ T; affine hull S = affine hull T\ \ (rel_interior S) \ (rel_interior T)" by (auto simp: rel_interior_def) lemma rel_interior_maximal: "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)" by (auto simp: rel_interior_def) lemma rel_interior: "rel_interior S = {x \ S. \T. open T \ x \ T \ T \ affine hull S \ S}" unfolding rel_interior_def[of S] openin_open[of "affine hull S"] apply auto proof - fix x T assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S" then have **: "x \ T \ affine hull S" using hull_inc by auto show "\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S" apply (rule_tac x = "T \ (affine hull S)" in exI) using * ** apply auto done qed lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)" by (auto simp: rel_interior) lemma mem_rel_interior_ball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_ball) apply (rule_tac x = "ball x e" in exI, simp) done lemma rel_interior_ball: "rel_interior S = {x \ S. \e. e > 0 \ ball x e \ affine hull S \ S}" using mem_rel_interior_ball [of _ S] by auto lemma mem_rel_interior_cball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_cball) apply (rule_tac x = "ball x e" in exI) apply (simp add: subset_trans [OF ball_subset_cball], auto) done lemma rel_interior_cball: "rel_interior S = {x \ S. \e. e > 0 \ cball x e \ affine hull S \ S}" using mem_rel_interior_cball [of _ S] by auto lemma rel_interior_empty [simp]: "rel_interior {} = {}" by (auto simp: rel_interior_def) lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" by (metis affine_hull_eq affine_sing) lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" apply (auto simp: rel_interior_ball) apply (rule_tac x=1 in exI, force) done lemma subset_rel_interior: fixes S T :: "'n::euclidean_space set" assumes "S \ T" and "affine hull S = affine hull T" shows "rel_interior S \ rel_interior T" using assms by (auto simp: rel_interior_def) lemma rel_interior_subset: "rel_interior S \ S" by (auto simp: rel_interior_def) lemma rel_interior_subset_closure: "rel_interior S \ closure S" using rel_interior_subset by (auto simp: closure_def) lemma interior_subset_rel_interior: "interior S \ rel_interior S" by (auto simp: rel_interior interior_def) lemma interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "aff_dim S = int(DIM('n))" shows "rel_interior S = interior S" proof - have "affine hull S = UNIV" using assms affine_hull_UNIV[of S] by auto then show ?thesis unfolding rel_interior interior_def by auto qed lemma rel_interior_interior: fixes S :: "'n::euclidean_space set" assumes "affine hull S = UNIV" shows "rel_interior S = interior S" using assms unfolding rel_interior interior_def by auto lemma rel_interior_open: fixes S :: "'n::euclidean_space set" assumes "open S" shows "rel_interior S = S" by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) -lemma interior_ball [simp]: "interior (ball x e) = ball x e" - by (simp add: interior_open) - lemma interior_rel_interior_gen: fixes S :: "'n::euclidean_space set" shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" by (metis interior_rel_interior low_dim_interior) lemma rel_interior_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_interior S = interior S" by (metis interior_rel_interior_gen) lemma affine_hull_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ affine hull S = UNIV" by (metis affine_hull_UNIV interior_rel_interior_gen) lemma rel_interior_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "rel_interior (affine hull S) = affine hull S" proof - have *: "rel_interior (affine hull S) \ affine hull S" using rel_interior_subset by auto { fix x assume x: "x \ affine hull S" define e :: real where "e = 1" then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto then have "x \ rel_interior (affine hull S)" using x rel_interior_ball[of "affine hull S"] by auto } then show ?thesis using * by auto qed lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" by (metis open_UNIV rel_interior_open) lemma rel_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto { fix y assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S" have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "x \ affine hull S" using assms hull_subset[of S] by auto moreover have "1 / e + - ((1 - e) / e) = 1" using \e > 0\ left_diff_distrib[of "1" "(1-e)" "1/e"] by auto ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ affine hull S" using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) done also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[THEN subsetD]) unfolding mem_ball using assms(3-5) ** apply auto done } then have "ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S" by auto moreover have "e * d > 0" using \e > 0\ \d > 0\ by simp moreover have c: "c \ S" using assms rel_interior_subset by auto moreover from c have "x - e *\<^sub>R (x - c) \ S" using convexD_alt[of S x c e] apply (simp add: algebra_simps) using assms apply auto done ultimately show ?thesis using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \e > 0\ by auto qed lemma interior_real_atLeast [simp]: fixes a :: real shows "interior {a..} = {a<..}" proof - { fix y assume "a < y" then have "y \ interior {a..}" apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {a..}" then obtain e where e: "e > 0" "cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreover from e have "y - e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y - e" by blast then have "a < y" using e by auto } ultimately show ?thesis by auto qed lemma continuous_ge_on_Ioo: assumes "continuous_on {c..d} g" "\x. x \ {c<.. g x \ a" "c < d" "x \ {c..d}" shows "g (x::real) \ (a::real)" proof- from assms(3) have "{c..d} = closure {c<.. (g -` {a..} \ {c..d})" by auto hence "closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono) also from assms(1) have "closed (g -` {a..} \ {c..d})" by (auto simp: continuous_on_closed_vimage) hence "closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp finally show ?thesis using \x \ {c..d}\ by auto qed lemma interior_real_atMost [simp]: fixes a :: real shows "interior {..a} = {.. y" then have "y \ interior {..a}" apply (simp add: mem_interior) apply (rule_tac x="(a-y)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {..a}" then obtain e where e: "e > 0" "cball y e \ {..a}" using mem_interior_cball[of y "{..a}"] by auto moreover from e have "y + e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y + e" by auto then have "a > y" using e by auto } ultimately show ?thesis by auto qed lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<.. {..b}" by auto also have "interior \ = {a<..} \ {.. = {a<.. {}" using assms unfolding set_eq_iff by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) then show ?thesis using interior_rel_interior_gen[of "cbox a b", symmetric] by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) qed lemma rel_interior_real_semiline [simp]: fixes a :: real shows "rel_interior {a..} = {a<..}" proof - have *: "{a<..} \ {}" unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"] by (auto split: if_split_asm) qed subsubsection \Relative open sets\ definition\<^marker>\tag important\ "rel_open S \ rel_interior S = S" lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" unfolding rel_open_def rel_interior_def apply auto using openin_subopen[of "top_of_set (affine hull S)" S] apply auto done lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" apply (simp add: rel_interior_def) apply (subst openin_subopen, blast) done lemma openin_set_rel_interior: "openin (top_of_set S) (rel_interior S)" by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) lemma affine_rel_open: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "rel_open S" unfolding rel_open_def using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] by metis lemma affine_closed: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "closed S" proof - { assume "S \ {}" then obtain L where L: "subspace L" "affine_parallel S L" using assms affine_parallel_subspace[of S] by auto then obtain a where a: "S = ((+) a ` L)" using affine_parallel_def[of L S] affine_parallel_commut by auto from L have "closed L" using closed_subspace by auto then have "closed S" using closed_translation a by auto } then show ?thesis by auto qed lemma closure_affine_hull: fixes S :: "'n::euclidean_space set" shows "closure S \ affine hull S" by (intro closure_minimal hull_subset affine_closed affine_affine_hull) lemma closed_affine_hull [iff]: fixes S :: "'n::euclidean_space set" shows "closed (affine hull S)" by (metis affine_affine_hull affine_closed) lemma closure_same_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "affine hull (closure S) = affine hull S" proof - have "affine hull (closure S) \ affine hull S" using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto moreover have "affine hull (closure S) \ affine hull S" using hull_mono[of "S" "closure S" "affine"] closure_subset by auto ultimately show ?thesis by auto qed lemma closure_aff_dim [simp]: fixes S :: "'n::euclidean_space set" shows "aff_dim (closure S) = aff_dim S" proof - have "aff_dim S \ aff_dim (closure S)" using aff_dim_subset closure_subset by auto moreover have "aff_dim (closure S) \ aff_dim (affine hull S)" using aff_dim_subset closure_affine_hull by blast moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto ultimately show ?thesis by auto qed lemma rel_interior_closure_convex_shrink: fixes S :: "_::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ closure S" and "e > 0" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto have "\y \ S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x], auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have zball: "z \ ball c d" using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (simp add: norm_minus_commute) (simp add: field_simps) have "x \ affine hull S" using closure_affine_hull assms by auto moreover have "y \ affine hull S" using \y \ S\ hull_subset[of S] by auto moreover have "c \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto ultimately have "z \ affine hull S" using z_def affine_affine_hull[of S] mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] assms by simp then have "z \ S" using d zball by auto obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d" using zball open_ball[of c d] openE[of "ball c d" z] by auto then have "ball z d1 \ affine hull S \ ball c d \ affine hull S" by auto then have "ball z d1 \ affine hull S \ S" using d by auto then have "z \ rel_interior S" using mem_rel_interior_ball using \d1 > 0\ \z \ S\ by auto then have "y - e *\<^sub>R (y - z) \ rel_interior S" using rel_interior_convex_shrink[of S z y e] assms \y \ S\ by auto then show ?thesis using * by auto qed lemma rel_interior_eq: "rel_interior s = s \ openin(top_of_set (affine hull s)) s" using rel_open rel_open_def by blast lemma rel_interior_openin: "openin(top_of_set (affine hull s)) s \ rel_interior s = s" by (simp add: rel_interior_eq) lemma rel_interior_affine: fixes S :: "'n::euclidean_space set" shows "affine S \ rel_interior S = S" using affine_rel_open rel_open_def by auto lemma rel_interior_eq_closure: fixes S :: "'n::euclidean_space set" shows "rel_interior S = closure S \ affine S" proof (cases "S = {}") case True then show ?thesis by auto next case False show ?thesis proof assume eq: "rel_interior S = closure S" have "S = {} \ S = affine hull S" apply (rule connected_clopen [THEN iffD1, rule_format]) apply (simp add: affine_imp_convex convex_connected) apply (rule conjI) apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) done with False have "affine hull S = S" by auto then show "affine S" by (metis affine_hull_eq) next assume "affine S" then show "rel_interior S = closure S" by (simp add: rel_interior_affine affine_closed) qed qed subsubsection\<^marker>\tag unimportant\\Relative interior preserves under linear transformations\ lemma rel_interior_translation_aux: fixes a :: "'n::euclidean_space" shows "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" proof - { fix x assume x: "x \ rel_interior S" then obtain T where "open T" "x \ T \ S" "T \ affine hull S \ S" using mem_rel_interior[of x S] by auto then have "open ((\x. a + x) ` T)" and "a + x \ ((\x. a + x) ` T) \ ((\x. a + x) ` S)" and "((\x. a + x) ` T) \ affine hull ((\x. a + x) ` S) \ (\x. a + x) ` S" using affine_hull_translation[of a S] open_translation[of T a] x by auto then have "a + x \ rel_interior ((\x. a + x) ` S)" using mem_rel_interior[of "a+x" "((\x. a + x) ` S)"] by auto } then show ?thesis by auto qed lemma rel_interior_translation: fixes a :: "'n::euclidean_space" shows "rel_interior ((\x. a + x) ` S) = (\x. a + x) ` rel_interior S" proof - have "(\x. (-a) + x) ` rel_interior ((\x. a + x) ` S) \ rel_interior S" using rel_interior_translation_aux[of "-a" "(\x. a + x) ` S"] translation_assoc[of "-a" "a"] by auto then have "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] by auto then show ?thesis using rel_interior_translation_aux[of a S] by auto qed lemma affine_hull_linear_image: assumes "bounded_linear f" shows "f ` (affine hull s) = affine hull f ` s" proof - interpret f: bounded_linear f by fact have "affine {x. f x \ affine hull f ` s}" unfolding affine_def by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) moreover have "affine {x. x \ f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) ultimately show ?thesis by (auto simp: hull_inc elim!: hull_induct) qed lemma rel_interior_injective_on_span_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and S :: "'m::euclidean_space set" assumes "bounded_linear f" and "inj_on f (span S)" shows "rel_interior (f ` S) = f ` (rel_interior S)" proof - { fix z assume z: "z \ rel_interior (f ` S)" then have "z \ f ` S" using rel_interior_subset[of "f ` S"] by auto then obtain x where x: "x \ S" "f x = z" by auto obtain e2 where e2: "e2 > 0" "cball z e2 \ affine hull (f ` S) \ (f ` S)" using z rel_interior_cball[of "f ` S"] by auto obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K" using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto define e1 where "e1 = 1 / K" then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x" using K pos_le_divide_eq[of e1] by auto define e where "e = e1 * e2" then have "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball x e \ affine hull S" then have h1: "f y \ affine hull (f ` S)" using affine_hull_linear_image[of f S] assms by auto from y have "norm (x-y) \ e1 * e2" using cball_def[of x e] dist_norm[of x y] e_def by auto moreover have "f x - f y = f (x - y)" using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto moreover have "e1 * norm (f (x-y)) \ norm (x - y)" using e1 by auto ultimately have "e1 * norm ((f x)-(f y)) \ e1 * e2" by auto then have "f y \ cball z e2" using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto then have "f y \ f ` S" using y e2 h1 by auto then have "y \ S" using assms y hull_subset[of S] affine_hull_subset_span inj_on_image_mem_iff [OF \inj_on f (span S)\] by (metis Int_iff span_superset subsetCE) } then have "z \ f ` (rel_interior S)" using mem_rel_interior_cball[of x S] \e > 0\ x by auto } moreover { fix x assume x: "x \ rel_interior S" then obtain e2 where e2: "e2 > 0" "cball x e2 \ affine hull S \ S" using rel_interior_cball[of S] by auto have "x \ S" using x rel_interior_subset by auto then have *: "f x \ f ` S" by auto have "\x\span S. f x = 0 \ x = 0" using assms subspace_span linear_conv_bounded_linear[of f] linear_injective_on_subspace_0[of f "span S"] by auto then obtain e1 where e1: "e1 > 0" "\x \ span S. e1 * norm x \ norm (f x)" using assms injective_imp_isometric[of "span S" f] subspace_span[of S] closed_subspace[of "span S"] by auto define e where "e = e1 * e2" hence "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball (f x) e \ affine hull (f ` S)" then have "y \ f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto then obtain xy where xy: "xy \ affine hull S" "f xy = y" by auto with y have "norm (f x - f xy) \ e1 * e2" using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto moreover have "f x - f xy = f (x - xy)" using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto moreover have *: "x - xy \ span S" using subspace_diff[of "span S" x xy] subspace_span \x \ S\ xy affine_hull_subset_span[of S] span_superset by auto moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))" using e1 by auto ultimately have "e1 * norm (x - xy) \ e1 * e2" by auto then have "xy \ cball x e2" using cball_def[of x e2] dist_norm[of x xy] e1 by auto then have "y \ f ` S" using xy e2 by auto } then have "f x \ rel_interior (f ` S)" using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \e > 0\ by auto } ultimately show ?thesis by auto qed lemma rel_interior_injective_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "bounded_linear f" and "inj f" shows "rel_interior (f ` S) = f ` (rel_interior S)" using assms rel_interior_injective_on_span_linear_image[of f S] subset_inj_on[of f "UNIV" "span S"] by auto subsection\<^marker>\tag unimportant\ \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (convex hull S)" proof (clarsimp simp: open_contains_cball convex_hull_explicit) fix T and u :: "'a\real" assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1" from assms[unfolded open_contains_cball] obtain b where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis have "b ` T \ {}" using obt by auto define i where "i = b ` T" let ?\ = "\y. \F. finite F \ F \ S \ (\u. (\x\F. 0 \ u x) \ sum u F = 1 \ (\v\F. u v *\<^sub>R v) = y)" let ?a = "\v\T. u v *\<^sub>R v" show "\e > 0. cball ?a e \ {y. ?\ y}" proof (intro exI subsetI conjI) show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` T\{}\] using b \T\S\ by auto next fix y assume "y \ cball ?a (Min i)" then have y: "norm (?a - y) \ Min i" unfolding dist_norm[symmetric] by auto { fix x assume "x \ T" then have "Min i \ b x" by (simp add: i_def obt(1)) then have "x + (y - ?a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto moreover have "x \ S" using \x\T\ \T\S\ by auto ultimately have "x + (y - ?a) \ S" using y b by blast } moreover have *: "inj_on (\v. v + (y - ?a)) T" unfolding inj_on_def by auto have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" unfolding sum.reindex[OF *] o_def using obt(4) by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) ultimately show "y \ {y. ?\ y}" proof (intro CollectI exI conjI) show "finite ((\v. v + (y - ?a)) ` T)" by (simp add: obt(1)) show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1" unfolding sum.reindex[OF *] o_def using obt(4) by auto qed (use obt(1, 3) in auto) qed qed lemma compact_convex_combinations: fixes S T :: "'a::real_normed_vector set" assumes "compact S" "compact T" shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}" proof - let ?X = "{0..1} \ S \ T" let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X" by force have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) with assms show ?thesis by (simp add: * compact_Times compact_continuous_image) qed lemma finite_imp_compact_convex_hull: fixes S :: "'a::real_normed_vector set" assumes "finite S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms show ?thesis proof (induct rule: finite_ne_induct) case (singleton x) show ?case by simp next case (insert x A) let ?f = "\(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" let ?T = "{0..1::real} \ (convex hull A)" have "continuous_on ?T ?f" unfolding split_def continuous_on by (intro ballI tendsto_intros) moreover have "compact ?T" by (intro compact_Times compact_Icc insert) ultimately have "compact (?f ` ?T)" by (rule compact_continuous_image) also have "?f ` ?T = convex hull (insert x A)" unfolding convex_hull_insert [OF \A \ {}\] apply safe apply (rule_tac x=a in exI, simp) apply (rule_tac x="1 - a" in exI, simp, fast) apply (rule_tac x="(u, b)" in image_eqI, simp_all) done finally show "compact (convex hull (insert x A))" . qed qed lemma compact_convex_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis using compact_empty by simp next case False then obtain w where "w \ S" by auto show ?thesis unfolding caratheodory[of S] proof (induct ("DIM('a) + 1")) case 0 have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}" using compact_empty by auto from 0 show ?case unfolding * by simp next case (Suc n) show ?case proof (cases "n = 0") case True have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "x \ S" proof (cases "card T = 0") case True then show ?thesis using T(4) unfolding card_0_eq[OF T(1)] by simp next case False then have "card T = Suc 0" using T(3) \n=0\ by auto then obtain a where "T = {a}" unfolding card_Suc_eq by auto then show ?thesis using T(2,4) by simp qed next fix x assume "x\S" then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="{x}" in exI) unfolding convex_hull_singleton apply auto done qed then show ?thesis using assms by simp next case False have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ {x. \T. finite T \ T \ S \ card T \ n \ x \ convex hull T}}" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T" by auto moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T" apply (rule convexD_alt) using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex] using obt(7) and hull_mono[of T "insert u T"] apply auto done ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="insert u T" in exI) apply (auto simp: card_insert_if) done next fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" proof (cases "card T = Suc n") case False then have "card T \ n" using T(3) by auto then show ?thesis apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using \w\S\ and T apply (auto intro!: exI[where x=T]) done next case True then obtain a u where au: "T = insert a u" "a\u" apply (drule_tac card_eq_SucD, auto) done show ?thesis proof (cases "u = {}") case True then have "x = a" using T(4)[unfolded au] by auto show ?thesis unfolding \x = a\ apply (rule_tac x=a in exI) apply (rule_tac x=a in exI) apply (rule_tac x=1 in exI) using T and \n \ 0\ unfolding au apply (auto intro!: exI[where x="{a}"]) done next case False obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using T(4)[unfolded au convex_hull_insert[OF False]] by auto have *: "1 - vx = ux" using obt(3) by auto show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=b in exI) apply (rule_tac x=vx in exI) using obt and T(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] apply (auto intro!: exI[where x=u]) done qed qed qed then show ?thesis using compact_convex_combinations[OF assms Suc] by simp qed qed qed subsection\<^marker>\tag unimportant\ \Extremal points of a simplex are some vertices\ lemma dist_increases_online: fixes a b d :: "'a::real_inner" assumes "d \ 0" shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" proof (cases "inner a d - inner b d > 0") case True then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" apply (rule_tac add_pos_pos) using assms apply auto done then show ?thesis apply (rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done next case False then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" apply (rule_tac add_pos_nonneg) using assms apply auto done then show ?thesis apply (rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done qed lemma norm_increases_online: fixes d :: "'a::real_inner" shows "d \ 0 \ norm (a + d) > norm a \ norm(a - d) > norm a" using dist_increases_online[of d a 0] unfolding dist_norm by auto lemma simplex_furthest_lt: fixes S :: "'a::real_inner set" assumes "finite S" shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))" using assms proof induct fix x S assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))" show "\xa\convex hull insert x S. xa \ insert x S \ (\y\convex hull insert x S. norm (xa - a) < norm (y - a))" proof (intro impI ballI, cases "S = {}") case False fix y assume y: "y \ convex hull insert x S" "y \ insert x S" obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto show "\z\convex hull insert x S. norm (y - a) < norm (z - a)" proof (cases "y \ convex hull S") case True then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)" using as(3)[THEN bspec[where x=y]] and y(2) by auto then show ?thesis apply (rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] apply auto done next case False show ?thesis using obt(3) proof (cases "u = 0", case_tac[!] "v = 0") assume "u = 0" "v \ 0" then have "y = b" using obt by auto then show ?thesis using False and obt(4) by auto next assume "u \ 0" "v = 0" then have "y = x" using obt by auto then show ?thesis using y(2) by auto next assume "u \ 0" "v \ 0" then obtain w where w: "w>0" "w b" proof assume "x = b" then have "y = b" unfolding obt(5) using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) then show False using obt(4) and False by simp qed then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto show ?thesis using dist_increases_online[OF *, of a y] proof (elim disjE) assume "dist a y < dist a (y + w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u + w \ 0" "v - w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto next assume "dist a y < dist a (y - w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u - w \ 0" "v + w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto qed qed auto qed qed auto qed (auto simp: assms) lemma simplex_furthest_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)" using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \finite S\], of a] unfolding dist_commute[of a] unfolding dist_norm by auto show ?thesis proof (cases "x \ S") case False then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto then show ?thesis using x(2)[THEN bspec[where x=y]] by auto next case True with x show ?thesis by auto qed qed lemma simplex_furthest_le_exists: fixes S :: "('a::real_inner) set" shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)" using simplex_furthest_le[of S] by (cases "S = {}") auto lemma simplex_extremal_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain u v where obt: "u \ convex hull S" "v \ convex hull S" "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)" using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) then show ?thesis proof (cases "u\S \ v\S", elim disjE) assume "u \ S" then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto then show ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto next assume "v \ S" then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto then show ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) by (auto simp: norm_minus_commute) qed auto qed lemma simplex_extremal_le_exists: fixes S :: "'a::real_inner set" shows "finite S \ x \ convex hull S \ y \ convex hull S \ \u\S. \v\S. norm (x - y) \ norm (u - v)" using convex_hull_empty simplex_extremal_le[of S] by(cases "S = {}") auto subsection \Closest point of a convex set is unique, with a continuous projection\ definition\<^marker>\tag important\ closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: assumes "closed S" and "S \ {}" - shows "closest_point S a \ S" + shows closest_point_in_set: "closest_point S a \ S" and "\y\S. dist a (closest_point S a) \ dist a y" unfolding closest_point_def apply(rule_tac[!] someI2_ex) apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) done -lemma closest_point_in_set: "closed S \ S \ {} \ closest_point S a \ S" - by (meson closest_point_exists) - lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x" using closest_point_exists[of S] by auto lemma closest_point_self: assumes "x \ S" shows "closest_point S x = x" unfolding closest_point_def apply (rule some1_equality, rule ex1I[of _ x]) using assms apply auto done lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S" using closest_point_in_set[of S x] closest_point_self[of x S] by auto lemma closer_points_lemma: assumes "inner y z > 0" shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y" proof - have z: "inner z z > 0" unfolding inner_gt_zero_iff using assms by auto have "norm (v *\<^sub>R z - y) < norm y" if "0 < v" and "v \ inner y z / inner z z" for v unfolding norm_lt using z assms that by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) then show ?thesis using assms z by (rule_tac x = "inner y z / inner z z" in exI) auto qed lemma closer_point_lemma: assumes "inner (y - x) (z - x) > 0" shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y" proof - obtain u where "u > 0" and u: "\v>0. v \ u \ norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" using closer_points_lemma[OF assms] by auto show ?thesis apply (rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and \u > 0\ unfolding dist_norm by (auto simp: norm_minus_commute field_simps) qed lemma any_closest_point_dot: assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" shows "inner (a - x) (y - x) \ 0" proof (rule ccontr) assume "\ ?thesis" then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \ S" using convexD_alt[OF assms(1,3,4), of u] using u by auto then show False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp: dist_commute algebra_simps) qed lemma any_closest_point_unique: fixes x :: "'a::real_inner" assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] unfolding norm_pths(1) and norm_le_square by (auto simp: algebra_simps) lemma closest_point_unique: assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z" shows "x = closest_point S a" using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] using closest_point_exists[OF assms(2)] and assms(3) by auto lemma closest_point_dot: assumes "convex S" "closed S" "x \ S" shows "inner (a - closest_point S a) (x - closest_point S a) \ 0" apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) using closest_point_exists[OF assms(2)] and assms(3) apply auto done lemma closest_point_lt: assumes "convex S" "closed S" "x \ S" "x \ closest_point S a" shows "dist a (closest_point S a) < dist a x" apply (rule ccontr) apply (rule_tac notE[OF assms(4)]) apply (rule closest_point_unique[OF assms(1-3), of a]) using closest_point_le[OF assms(2), of _ a] apply fastforce done lemma setdist_closest_point: "\closed S; S \ {}\ \ setdist {a} S = dist a (closest_point S a)" apply (rule setdist_unique) using closest_point_le apply (auto simp: closest_point_in_set) done lemma closest_point_lipschitz: assumes "convex S" and "closed S" "S \ {}" shows "dist (closest_point S x) (closest_point S y) \ dist x y" proof - have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0" and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0" apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) using closest_point_exists[OF assms(2-3)] apply auto done then show ?thesis unfolding dist_norm and norm_le using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous (at x) (closest_point S)" unfolding continuous_at_eps_delta using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto lemma continuous_on_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous_on t (closest_point S)" by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) proposition closest_point_in_rel_interior: assumes "closed S" "S \ {}" and x: "x \ affine hull S" shows "closest_point S x \ rel_interior S \ x \ rel_interior S" proof (cases "x \ S") case True then show ?thesis by (simp add: closest_point_self) next case False then have "False" if asm: "closest_point S x \ rel_interior S" proof - obtain e where "e > 0" and clox: "closest_point S x \ S" and e: "cball (closest_point S x) e \ affine hull S \ S" using asm mem_rel_interior_cball by blast then have clo_notx: "closest_point S x \ x" using \x \ S\ by auto define y where "y \ closest_point S x - (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" by (simp add: y_def algebra_simps) then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" by simp also have "\ < norm(x - closest_point S x)" using clo_notx \e > 0\ by (auto simp: mult_less_cancel_right2 field_split_simps) finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . have "y \ affine hull S" unfolding y_def by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) moreover have "dist (closest_point S x) y \ e" using \e > 0\ by (auto simp: y_def min_mult_distrib_right) ultimately have "y \ S" using subsetD [OF e] by simp then have "dist x (closest_point S x) \ dist x y" by (simp add: closest_point_le \closed S\) with no_less show False by (simp add: dist_norm) qed moreover have "x \ rel_interior S" using rel_interior_subset False by blast ultimately show ?thesis by blast qed subsubsection\<^marker>\tag unimportant\ \Various point-to-set separating/supporting hyperplane theorems\ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "S \ {}" and "z \ S" shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)" proof - obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2-3)]) show ?thesis proof (intro exI bexI conjI ballI) show "(y - z) \ z < (y - z) \ y" by (metis \y \ S\ assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x proof (rule ccontr) have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" using assms(1)[unfolded convex_alt] and y and \x\S\ and \y\S\ by auto assume "\ (y - z) \ y \ (y - z) \ x" then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] by (auto simp: inner_diff) then show False using *[of v] by (auto simp: dist_commute algebra_simps) qed qed (use \y \ S\ in auto) qed lemma separating_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "z \ S" shows "\a b. inner a z < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True then show ?thesis by (simp add: gt_ex) next case False obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2) False]) show ?thesis proof (intro exI conjI ballI) show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" using \y\S\ \z\S\ by auto next fix x assume "x \ S" have "False" if *: "0 < inner (z - y) (x - y)" proof - obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" using * closer_point_lemma by blast then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\] using \x\S\ \y\S\ by (auto simp: dist_commute algebra_simps) qed moreover have "0 < (norm (y - z))\<^sup>2" using \y\S\ \z\S\ by auto then have "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) qed qed lemma separating_hyperplane_closed_0: assumes "convex (S::('a::euclidean_space) set)" and "closed S" and "0 \ S" shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True have "(SOME i. i\Basis) \ (0::'a)" by (metis Basis_zero SOME_Basis) then show ?thesis using True zero_less_one by blast next case False then show ?thesis using False using separating_hyperplane_closed_point[OF assms] by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) qed subsubsection\<^marker>\tag unimportant\ \Now set-to-set for closed/compact sets\ lemma separating_hyperplane_closed_compact: fixes S :: "'a::euclidean_space set" assumes "convex S" and "closed S" and "convex T" and "compact T" and "T \ {}" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof (cases "S = {}") case True obtain b where b: "b > 0" "\x\T. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto obtain z :: 'a where z: "norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto then have "z \ T" using b(2)[THEN bspec[where x=z]] by auto then obtain a b where ab: "inner a z < b" "\x\T. b < inner a x" using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto then show ?thesis using True by auto next case False then obtain y where "y \ S" by auto obtain a b where "0 < b" "\x \ (\x\ S. \y \ T. {x - y}). b < inner a x" using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] using closed_compact_differences[OF assms(2,4)] using assms(6) by auto then have ab: "\x\S. \y\T. b + inner a y < inner a x" apply - apply rule apply rule apply (erule_tac x="x - y" in ballE) apply (auto simp: inner_diff) done define k where "k = (SUP x\T. a \ x)" show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-(k + b / 2)" in exI) apply (intro conjI ballI) unfolding inner_minus_left and neg_less_iff_less proof - fix x assume "x \ T" then have "inner a x - b / 2 < k" unfolding k_def proof (subst less_cSUP_iff) show "T \ {}" by fact show "bdd_above ((\) a ` T)" using ab[rule_format, of y] \y \ S\ by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) qed (auto intro!: bexI[of _ x] \0) then show "inner a x < k + b / 2" by auto next fix x assume "x \ S" then have "k \ inner a x - b" unfolding k_def apply (rule_tac cSUP_least) using assms(5) using ab[THEN bspec[where x=x]] apply auto done then show "k + b / 2 < inner a x" using \0 < b\ by auto qed qed lemma separating_hyperplane_compact_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "S \ {}" and "convex T" and "closed T" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof - obtain a b where "(\x\T. inner a x < b) \ (\x\S. b < inner a x)" using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto then show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-b" in exI, auto) done qed subsubsection\<^marker>\tag unimportant\ \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: assumes "convex S" "(0::'a::euclidean_space) \ S" shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)" proof - let ?k = "\c. {x::'a. 0 \ inner c x}" have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f proof - obtain c where c: "f = ?k ` c" "c \ S" "finite c" using finite_subset_image[OF as(2,1)] by auto then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) using subset_hull[of convex, OF assms(1), symmetric, of c] by force then have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)" apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) using hull_subset[of c convex] unfolding subset_eq and inner_scaleR by (auto simp: inner_commute del: ballE elim!: ballE) then show "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball sphere_def dist_norm by auto qed have "frontier (cball 0 1) \ (\(?k ` S)) \ {}" apply (rule compact_imp_fip) apply (rule compact_frontier[OF compact_cball]) using * closed_halfspace_ge by auto then obtain x where "norm x = 1" "\y\S. x\?k y" unfolding frontier_cball dist_norm sphere_def by auto then show ?thesis by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) qed lemma separating_hyperplane_sets: fixes S T :: "'a::euclidean_space set" assumes "convex S" and "convex T" and "S \ {}" and "T \ {}" and "S \ T = {}" shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)" proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x" using assms(3-5) by force then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x" by (force simp: inner_diff) then have bdd: "bdd_above (((\) a)`S)" using \T \ {}\ by (auto intro: bdd_aboveI2[OF *]) show ?thesis using \a\0\ by (intro exI[of _ a] exI[of _ "SUP x\S. a \ x"]) (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \S \ {}\ *) qed subsection\<^marker>\tag unimportant\ \More convexity generalities\ lemma convex_closure [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (closure S)" apply (rule convexI) apply (unfold closure_sequential, elim exE) apply (rule_tac x="\n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) apply (rule,rule) apply (rule convexD [OF assms]) apply (auto del: tendsto_const intro!: tendsto_intros) done lemma convex_interior [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (interior S)" unfolding convex_alt Ball_def mem_interior proof clarify fix x y u assume u: "0 \ u" "u \ (1::real)" fix e d assume ed: "ball x e \ S" "ball y d \ S" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S" proof (intro exI conjI subsetI) fix z assume "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S" apply (rule_tac assms[unfolded convex_alt, rule_format]) using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm apply (auto simp: algebra_simps) done then show "z \ S" using u by (auto simp: algebra_simps) qed(insert u ed(3-4), auto) qed lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}" using hull_subset[of S convex] convex_hull_empty by auto subsection\<^marker>\tag unimportant\ \Convex set as intersection of halfspaces\ lemma convex_halfspace_intersection: fixes s :: "('a::euclidean_space) set" assumes "closed s" "convex s" shows "s = \{h. s \ h \ (\a b. h = {x. inner a x \ b})}" apply (rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply (rule,rule,erule conjE) proof - fix x assume "\xa. s \ xa \ (\a b. xa = {x. inner a x \ b}) \ x \ xa" then have "\a b. s \ {x. inner a x \ b} \ x \ {x. inner a x \ b}" by blast then show "x \ s" apply (rule_tac ccontr) apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) apply (erule exE)+ apply (erule_tac x="-a" in allE) apply (erule_tac x="-b" in allE, auto) done qed auto subsection\<^marker>\tag unimportant\ \Convexity of general and special intervals\ lemma is_interval_convex: fixes S :: "'a::euclidean_space set" assumes "is_interval S" shows "convex S" proof (rule convexI) fix x y and u v :: real assume as: "x \ S" "y \ S" "0 \ u" "0 \ v" "u + v = 1" then have *: "u = 1 - v" "1 - v \ 0" and **: "v = 1 - u" "1 - u \ 0" by auto { fix a b assume "\ b \ u * a + v * b" then have "u * a < (1 - v) * b" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) *(2) apply (rule_tac mult_left_less_imp_less[of "1 - v"]) apply (auto simp: field_simps) done then have "a \ u * a + v * b" unfolding * using as(4) by (auto simp: field_simps intro!:mult_right_mono) } moreover { fix a b assume "\ u * a + v * b \ a" then have "v * b > (1 - u) * a" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) apply (rule_tac mult_left_less_imp_less) apply (auto simp: field_simps) done then have "u * a + v * b \ b" unfolding ** using **(2) as(3) by (auto simp: field_simps intro!:mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ S" apply - apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) DIM_positive[where 'a='a] apply (auto simp: inner_simps) done qed lemma is_interval_connected: fixes S :: "'a::euclidean_space set" shows "is_interval S \ connected S" using is_interval_convex convex_connected by auto lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" apply (rule_tac[!] is_interval_convex)+ using is_interval_box is_interval_cbox apply auto done text\A non-singleton connected set is perfect (i.e. has no isolated points). \ lemma connected_imp_perfect: fixes a :: "'a::metric_space" assumes "connected S" "a \ S" and S: "\x. S \ {x}" shows "a islimpt S" proof - have False if "a \ T" "open T" "\y. \y \ S; y \ T\ \ y = a" for T proof - obtain e where "e > 0" and e: "cball a e \ T" using \open T\ \a \ T\ by (auto simp: open_contains_cball) have "openin (top_of_set S) {a}" unfolding openin_open using that \a \ S\ by blast moreover have "closedin (top_of_set S) {a}" by (simp add: assms) ultimately show "False" using \connected S\ connected_clopen S by blast qed then show ?thesis unfolding islimpt_def by blast qed lemma connected_imp_perfect_aff_dim: "\connected S; aff_dim S \ 0; a \ S\ \ a islimpt S" using aff_dim_sing connected_imp_perfect by blast subsection\<^marker>\tag unimportant\ \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ lemma mem_is_interval_1_I: fixes a b c::real assumes "is_interval S" assumes "a \ S" "c \ S" assumes "a \ b" "b \ c" shows "b \ S" using assms is_interval_1 by blast lemma is_interval_connected_1: fixes s :: "real set" shows "is_interval s \ connected s" apply rule apply (rule is_interval_connected, assumption) unfolding is_interval_1 apply rule apply rule apply rule apply rule apply (erule conjE) apply (rule ccontr) proof - fix a b x assume as: "connected s" "a \ s" "b \ s" "a \ x" "x \ b" "x \ s" then have *: "a < x" "x < b" unfolding not_le [symmetric] by auto let ?halfl = "{.. s" with \x \ s\ have "x \ y" by auto then have "y \ ?halfr \ ?halfl" by auto } moreover have "a \ ?halfl" "b \ ?halfr" using * by auto then have "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto ultimately show False apply (rule_tac notE[OF as(1)[unfolded connected_def]]) apply (rule_tac x = ?halfl in exI) apply (rule_tac x = ?halfr in exI, rule) apply (rule open_lessThan, rule) apply (rule open_greaterThan, auto) done qed lemma is_interval_convex_1: fixes s :: "real set" shows "is_interval s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_compact_interval_1: "connected S \ compact S \ (\a b. S = {a..b::real})" by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) lemma connected_convex_1: fixes s :: "real set" shows "connected s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_convex_1_gen: fixes s :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected s \ convex s" proof - obtain f:: "'a \ real" where linf: "linear f" and "inj f" using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, where 'a='a and 'b=real] unfolding Euclidean_Space.dim_UNIV by (auto simp: assms) then have "f -` (f ` s) = s" by (simp add: inj_vimage_image_eq) then show ?thesis by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) qed lemma [simp]: fixes r s::real shows is_interval_io: "is_interval {..\tag unimportant\ \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f a)\k \ y" "y \ (f b)\k" shows "\x\{a..b}. (f x)\k = y" proof - have "f a \ f ` cbox a b" "f b \ f ` cbox a b" apply (rule_tac[!] imageI) using assms(1) apply auto done then show ?thesis using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] by (simp add: connected_continuous_image assms) qed lemma ivt_increasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f b)\k \ y" and "y \ (f a)\k" shows "\x\{a..b}. (f x)\k = y" apply (subst neg_equal_iff_equal[symmetric]) using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] using assms using continuous_on_minus apply auto done lemma ivt_decreasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) subsection\<^marker>\tag unimportant\ \A bound within an interval\ lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" shows "convex hull {x, y} = cbox x y" proof (rule hull_unique) show "{x, y} \ cbox x y" using \x \ y\ by auto show "convex (cbox x y)" by (rule convex_box) next fix S assume "{x, y} \ S" and "convex S" then show "cbox x y \ S" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed lemma unit_interval_convex_hull: "cbox (0::'a::euclidean_space) One = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") proof - have One[simp]: "\i. i \ Basis \ One \ i = 1" by (simp add: inner_sum_left sum.If_cases inner_Basis) have "?int = {x. \i\Basis. x \ i \ cbox 0 1}" by (auto simp: cbox_def) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` cbox 0 1)" by (simp only: box_eq_set_sum_Basis) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" by (simp add: convex_hull_linear_image) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" by (simp only: convex_hull_set_sum) also have "\ = convex hull {x. \i\Basis. x\i \ {0, 1}}" by (simp only: box_eq_set_sum_Basis) also have "convex hull {x. \i\Basis. x\i \ {0, 1}} = convex hull ?points" by simp finally show ?thesis . qed text \And this is a finite set of vertices.\ lemma unit_cube_convex_hull: obtains S :: "'a::euclidean_space set" where "finite S" and "cbox 0 (\Basis) = convex hull S" proof show "finite {x::'a. \i\Basis. x \ i = 0 \ x \ i = 1}" proof (rule finite_subset, clarify) show "finite ((\S. \i\Basis. (if i \ S then 1 else 0) *\<^sub>R i) ` Pow Basis)" using finite_Basis by blast fix x :: 'a assume as: "\i\Basis. x \ i = 0 \ x \ i = 1" show "x \ (\S. \i\Basis. (if i\S then 1 else 0) *\<^sub>R i) ` Pow Basis" apply (rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) using as apply (subst euclidean_eq_iff, auto) done qed show "cbox 0 One = convex hull {x. \i\Basis. x \ i = 0 \ x \ i = 1}" using unit_interval_convex_hull by blast qed text \Hence any cube (could do any nonempty interval).\ lemma cube_convex_hull: assumes "d > 0" obtains S :: "'a::euclidean_space set" where "finite S" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull S" proof - let ?d = "(\i\Basis. d *\<^sub>R i)::'a" have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" proof (intro set_eqI iffI) fix y assume "y \ cbox (x - ?d) (x + ?d)" then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" using assms by (simp add: mem_box inner_simps) (simp add: field_simps) with \0 < d\ show "y \ (\y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) next fix y assume "y \ (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" then obtain z where z: "z \ cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" by auto then show "y \ cbox (x - ?d) (x + ?d)" using z assms by (auto simp: mem_box inner_simps) qed obtain S where "finite S" "cbox 0 (\Basis::'a) = convex hull S" using unit_cube_convex_hull by auto then show ?thesis by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed subsection\<^marker>\tag unimportant\\Representation of any interval as a finite convex hull\ lemma image_stretch_interval: "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = (if (cbox a b) = {} then {} else cbox (\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k::'a) (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k))" proof cases assume *: "cbox a b \ {}" show ?thesis unfolding box_ne_empty if_not_P[OF *] apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) apply (subst choice_Basis_iff[symmetric]) proof (intro allI ball_cong refl) fix x i :: 'a assume "i \ Basis" with * have a_le_b: "a \ i \ b \ i" unfolding box_ne_empty by auto show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" proof (cases "m i = 0") case True with a_le_b show ?thesis by auto next case False then have *: "\a b. a = m i * b \ b = a / m i" by (auto simp: field_simps) from False have "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) with False show ?thesis using a_le_b unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps) qed qed qed simp lemma interval_image_stretch_interval: "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" unfolding image_stretch_interval by auto lemma cbox_translation: "cbox (c + a) (c + b) = image (\x. c + x) (cbox a b)" using image_affinity_cbox [of 1 c a b] using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] by (auto simp: inner_left_distrib add.commute) lemma cbox_image_unit_interval: fixes a :: "'a::euclidean_space" assumes "cbox a b \ {}" shows "cbox a b = (+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` cbox 0 One" using assms apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) done lemma closed_interval_as_convex_hull: fixes a :: "'a::euclidean_space" obtains S where "finite S" "cbox a b = convex hull S" proof (cases "cbox a b = {}") case True with convex_hull_empty that show ?thesis by blast next case False obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" by (blast intro: unit_cube_convex_hull) have lin: "linear (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k)" by (rule linear_compose_sum) (auto simp: algebra_simps linearI) have "finite ((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` S)" by (rule finite_imageI \finite S\)+ then show ?thesis apply (rule that) apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) done qed subsection\<^marker>\tag unimportant\ \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: fixes S :: "('a::real_normed_vector) set" assumes "open S" and "convex_on S f" and "\x\S. \f x\ \ b" shows "continuous_on S f" apply (rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof (rule,rule,rule) fix x and e :: real assume "x \ S" "e > 0" define B where "B = \b\ + 1" then have B: "0 < B""\x. x\S \ \f x\ \ B" using assms(3) by auto obtain k where "k > 0" and k: "cball x k \ S" using \x \ S\ assms(1) open_contains_cball_eq by blast show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" proof (intro exI conjI allI impI) fix y assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" show "\f y - f x\ < e" proof (cases "y = x") case False define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ by (auto simp:field_simps) have "y \ S" apply (rule k[THEN subsetD]) unfolding mem_cball dist_norm apply (rule order_trans[of _ "2 * norm (x - y)"]) using as by (auto simp: field_simps norm_minus_commute) { define w where "w = x + t *\<^sub>R (y - x)" have "w \ S" using \k>0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = 0" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have 2: "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) have "f y - f x \ (f w - f x) / t" using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] using \0 < t\ \2 < t\ and \x \ S\ \w \ S\ by (auto simp:field_simps) also have "... < e" using B(2)[OF \w\S\] and B(2)[OF \x\S\] 2 \t > 0\ by (auto simp: field_simps) finally have th1: "f y - f x < e" . } moreover { define w where "w = x - t *\<^sub>R (y - x)" have "w \ S" using \k > 0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = x" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) then have *: "(f w - f y) / t < e" using B(2)[OF \w\S\] and B(2)[OF \y\S\] using \t > 0\ by (auto simp:field_simps) have "f x \ 1 / (1 + t) * f w + (t / (1 + t)) * f y" using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] using \0 < t\ \2 < t\ and \y \ S\ \w \ S\ by (auto simp:field_simps) also have "\ = (f w + t * f y) / (1 + t)" using \t > 0\ by (simp add: add_divide_distrib) also have "\ < e + f y" using \t > 0\ * \e > 0\ by (auto simp: field_simps) finally have "f x - f y < e" by auto } ultimately show ?thesis by auto qed (insert \0, auto) qed (insert \0 \0 \0, auto simp: field_simps) qed subsection\<^marker>\tag unimportant\ \Upper bound on a ball implies upper and lower bounds\ lemma convex_bounds_lemma: fixes x :: "'a::real_normed_vector" assumes "convex_on (cball x e) f" and "\y \ cball x e. f y \ b" shows "\y \ cball x e. \f y\ \ b + 2 * \f x\" apply rule proof (cases "0 \ e") case True fix y assume y: "y \ cball x e" define z where "z = 2 *\<^sub>R x - y" have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp: algebra_simps) then show "\f y\ \ b + 2 * \f x\" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by (auto simp:field_simps) next case False fix y assume "y \ cball x e" then have "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) then show "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed subsubsection\<^marker>\tag unimportant\ \Hence a convex function on an open set is continuous\ lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" by auto lemma convex_on_continuous: assumes "open (s::('a::euclidean_space) set)" "convex_on s f" shows "continuous_on s f" unfolding continuous_on_eq_continuous_at[OF assms(1)] proof note dimge1 = DIM_positive[where 'a='a] fix x assume "x \ s" then obtain e where e: "cball x e \ s" "e > 0" using assms(1) unfolding open_contains_cball by auto define d where "d = e / real DIM('a)" have "0 < d" unfolding d_def using \e > 0\ dimge1 by auto let ?d = "(\i\Basis. d *\<^sub>R i)::'a" obtain c where c: "finite c" and c1: "convex hull c \ cball x e" and c2: "cball x d \ convex hull c" proof define c where "c = (\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d})" show "finite c" unfolding c_def by (simp add: finite_set_sum) have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" unfolding box_eq_set_sum_Basis unfolding c_def convex_hull_set_sum apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_add_left) apply (rule sum.cong [OF refl]) apply (rule image_cong [OF _ refl]) apply (rule convex_hull_eq_real_cbox) apply (cut_tac \0 < d\, simp) done then have 2: "convex hull c = {a. \i\Basis. a \ i \ cball (x \ i) d}" by (simp add: dist_norm abs_le_iff algebra_simps) show "cball x d \ convex hull c" unfolding 2 by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) have e': "e = (\(i::'a)\Basis. d)" by (simp add: d_def) show "convex hull c \ cball x e" unfolding 2 apply clarsimp apply (subst euclidean_dist_l2) apply (rule order_trans [OF L2_set_le_sum]) apply (rule zero_le_dist) unfolding e' apply (rule sum_mono, simp) done qed define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" apply(rule convex_on_subset[OF assms(2)]) apply(rule subset_trans[OF c1 e(1)]) done then have k: "\y\convex hull c. f y \ k" apply (rule_tac convex_on_convex_hull_bound, assumption) by (simp add: k_def c) have "e \ e * real DIM('a)" using e(2) real_of_nat_ge_one_iff by auto then have "d \ e" by (simp add: d_def field_split_simps) then have dsube: "cball x d \ cball x e" by (rule subset_cball) have conv: "convex_on (cball x d) f" using \convex_on (convex hull c) f\ c2 convex_on_subset by blast then have "\y\cball x d. \f y\ \ k + 2 * \f x\" by (rule convex_bounds_lemma) (use c2 k in blast) then have "continuous_on (ball x d) f" apply (rule_tac convex_on_bounded_continuous) apply (rule open_ball, rule convex_on_subset[OF conv]) apply (rule ball_subset_cball, force) done then show "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using \d > 0\ by auto qed end diff --git a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy @@ -1,5080 +1,5080 @@ (* Title: HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Huge cleanup by LCP *) theory Equivalence_Lebesgue_Henstock_Integration imports Lebesgue_Measure Henstock_Kurzweil_Integration Complete_Measure Set_Integral Homeomorphism Cartesian_Euclidean_Space begin lemma le_left_mono: "x \ y \ y \ a \ x \ (a::'a::preorder)" by (auto intro: order_trans) lemma ball_trans: assumes "y \ ball z q" "r + q \ s" shows "ball y r \ ball z s" using assms by metric lemma has_integral_implies_lebesgue_measurable_cbox: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) (cbox x y)" shows "f \ lebesgue_on (cbox x y) \\<^sub>M borel" proof (rule cld_measure.borel_measurable_cld) let ?L = "lebesgue_on (cbox x y)" let ?\ = "emeasure ?L" let ?\' = "outer_measure_of ?L" interpret L: finite_measure ?L proof show "?\ (space ?L) \ \" by (simp add: emeasure_restrict_space space_restrict_space emeasure_lborel_cbox_eq) qed show "cld_measure ?L" proof fix B A assume "B \ A" "A \ null_sets ?L" then show "B \ sets ?L" using null_sets_completion_subset[OF \B \ A\, of lborel] by (auto simp add: null_sets_restrict_space sets_restrict_space_iff intro: ) next fix A assume "A \ space ?L" "\B. B \ sets ?L \ ?\ B < \ \ A \ B \ sets ?L" from this(1) this(2)[of "space ?L"] show "A \ sets ?L" by (auto simp: Int_absorb2 less_top[symmetric]) qed auto then interpret cld_measure ?L . have content_eq_L: "A \ sets borel \ A \ cbox x y \ content A = measure ?L A" for A by (subst measure_restrict_space) (auto simp: measure_def) fix E and a b :: real assume "E \ sets ?L" "a < b" "0 < ?\ E" "?\ E < \" then obtain M :: real where "?\ E = M" "0 < M" by (cases "?\ E") auto define e where "e = M / (4 + 2 / (b - a))" from \a < b\ \0 have "0 < e" by (auto intro!: divide_pos_pos simp: field_simps e_def) have "e < M / (3 + 2 / (b - a))" using \a < b\ \0 < M\ unfolding e_def by (intro divide_strict_left_mono add_strict_right_mono mult_pos_pos) (auto simp: field_simps) then have "2 * e < (b - a) * (M - e * 3)" using \0 \0 < e\ \a < b\ by (simp add: field_simps) have e_less_M: "e < M / 1" unfolding e_def using \a < b\ \0 by (intro divide_strict_left_mono) (auto simp: field_simps) obtain d where "gauge d" and integral_f: "\p. p tagged_division_of cbox x y \ d fine p \ norm ((\(x,k) \ p. content k *\<^sub>R f x) - I) < e" using \0 f unfolding has_integral by auto define C where "C X m = X \ {x. ball x (1/Suc m) \ d x}" for X m have "incseq (C X)" for X unfolding C_def [abs_def] by (intro monoI Collect_mono conj_mono imp_refl le_left_mono subset_ball divide_left_mono Int_mono) auto { fix X assume "X \ space ?L" and eq: "?\' X = ?\ E" have "(SUP m. outer_measure_of ?L (C X m)) = outer_measure_of ?L (\m. C X m)" using \X \ space ?L\ by (intro SUP_outer_measure_of_incseq \incseq (C X)\) (auto simp: C_def) also have "(\m. C X m) = X" proof - { fix x obtain e where "0 < e" "ball x e \ d x" using gaugeD[OF \gauge d\, of x] unfolding open_contains_ball by auto moreover obtain n where "1 / (1 + real n) < e" using reals_Archimedean[OF \0] by (auto simp: inverse_eq_divide) then have "ball x (1 / (1 + real n)) \ ball x e" by (intro subset_ball) auto ultimately have "\n. ball x (1 / (1 + real n)) \ d x" by blast } then show ?thesis by (auto simp: C_def) qed finally have "(SUP m. outer_measure_of ?L (C X m)) = ?\ E" using eq by auto also have "\ > M - e" using \0 < M\ \?\ E = M\ \0 by (auto intro!: ennreal_lessI) finally have "\m. M - e < outer_measure_of ?L (C X m)" unfolding less_SUP_iff by auto } note C = this let ?E = "{x\E. f x \ a}" and ?F = "{x\E. b \ f x}" have "\ (?\' ?E = ?\ E \ ?\' ?F = ?\ E)" proof assume eq: "?\' ?E = ?\ E \ ?\' ?F = ?\ E" with C[of ?E] C[of ?F] \E \ sets ?L\[THEN sets.sets_into_space] obtain ma mb where "M - e < outer_measure_of ?L (C ?E ma)" "M - e < outer_measure_of ?L (C ?F mb)" by auto moreover define m where "m = max ma mb" ultimately have M_minus_e: "M - e < outer_measure_of ?L (C ?E m)" "M - e < outer_measure_of ?L (C ?F m)" using incseqD[OF \incseq (C ?E)\, of ma m, THEN outer_measure_of_mono] incseqD[OF \incseq (C ?F)\, of mb m, THEN outer_measure_of_mono] by (auto intro: less_le_trans) define d' where "d' x = d x \ ball x (1 / (3 * Suc m))" for x have "gauge d'" unfolding d'_def by (intro gauge_Int \gauge d\ gauge_ball) auto then obtain p where p: "p tagged_division_of cbox x y" "d' fine p" by (rule fine_division_exists) then have "d fine p" unfolding d'_def[abs_def] fine_def by auto define s where "s = {(x::'a, k). k \ (C ?E m) \ {} \ k \ (C ?F m) \ {}}" define T where "T E k = (SOME x. x \ k \ C E m)" for E k let ?A = "(\(x, k). (T ?E k, k)) ` (p \ s) \ (p - s)" let ?B = "(\(x, k). (T ?F k, k)) ` (p \ s) \ (p - s)" { fix X assume X_eq: "X = ?E \ X = ?F" let ?T = "(\(x, k). (T X k, k))" let ?p = "?T ` (p \ s) \ (p - s)" have in_s: "(x, k) \ s \ T X k \ k \ C X m" for x k using someI_ex[of "\x. x \ k \ C X m"] X_eq unfolding ex_in_conv by (auto simp: T_def s_def) { fix x k assume "(x, k) \ p" "(x, k) \ s" have k: "k \ ball x (1 / (3 * Suc m))" using \d' fine p\[THEN fineD, OF \(x, k) \ p\] by (auto simp: d'_def) then have "x \ ball (T X k) (1 / (3 * Suc m))" using in_s[OF \(x, k) \ s\] by (auto simp: C_def subset_eq dist_commute) then have "ball x (1 / (3 * Suc m)) \ ball (T X k) (1 / Suc m)" by (rule ball_trans) (auto simp: field_split_simps) with k in_s[OF \(x, k) \ s\] have "k \ d (T X k)" by (auto simp: C_def) } then have "d fine ?p" using \d fine p\ by (auto intro!: fineI) moreover have "?p tagged_division_of cbox x y" proof (rule tagged_division_ofI) show "finite ?p" using p(1) by auto next fix z k assume *: "(z, k) \ ?p" then consider "(z, k) \ p" "(z, k) \ s" | x' where "(x', k) \ p" "(x', k) \ s" "z = T X k" by (auto simp: T_def) then have "z \ k \ k \ cbox x y \ (\a b. k = cbox a b)" using p(1) by cases (auto dest: in_s) then show "z \ k" "k \ cbox x y" "\a b. k = cbox a b" by auto next fix z k z' k' assume "(z, k) \ ?p" "(z', k') \ ?p" "(z, k) \ (z', k')" with tagged_division_ofD(5)[OF p(1), of _ k _ k'] show "interior k \ interior k' = {}" by (auto simp: T_def dest: in_s) next have "{k. \x. (x, k) \ ?p} = {k. \x. (x, k) \ p}" by (auto simp: T_def image_iff Bex_def) then show "\{k. \x. (x, k) \ ?p} = cbox x y" using p(1) by auto qed ultimately have I: "norm ((\(x,k) \ ?p. content k *\<^sub>R f x) - I) < e" using integral_f by auto have "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) + (\(x,k) \ p - s. content k *\<^sub>R f x)" using p(1)[THEN tagged_division_ofD(1)] by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def) also have "(\(x,k) \ ?T ` (p \ s). content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k))" proof (subst sum.reindex_nontrivial, safe) fix x1 x2 k assume 1: "(x1, k) \ p" "(x1, k) \ s" and 2: "(x2, k) \ p" "(x2, k) \ s" and eq: "content k *\<^sub>R f (T X k) \ 0" with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k] show "x1 = x2" by (auto simp: content_eq_0_interior) qed (use p in \auto intro!: sum.cong\) finally have eq: "(\(x,k) \ ?p. content k *\<^sub>R f x) = (\(x,k) \ p \ s. content k *\<^sub>R f (T X k)) + (\(x,k) \ p - s. content k *\<^sub>R f x)" . have in_T: "(x, k) \ s \ T X k \ X" for x k using in_s[of x k] by (auto simp: C_def) note I eq in_T } note parts = this have p_in_L: "(x, k) \ p \ k \ sets ?L" for x k using tagged_division_ofD(3, 4)[OF p(1), of x k] by (auto simp: sets_restrict_space) have [simp]: "finite p" using tagged_division_ofD(1)[OF p(1)] . have "(M - 3*e) * (b - a) \ (\(x,k) \ p \ s. content k) * (b - a)" proof (intro mult_right_mono) have fin: "?\ (E \ \{k\snd`p. k \ C X m = {}}) < \" for X using \?\ E < \\ by (rule le_less_trans[rotated]) (auto intro!: emeasure_mono \E \ sets ?L\) have sets: "(E \ \{k\snd`p. k \ C X m = {}}) \ sets ?L" for X using tagged_division_ofD(1)[OF p(1)] by (intro sets.Diff \E \ sets ?L\ sets.finite_Union sets.Int) (auto intro: p_in_L) { fix X assume "X \ E" "M - e < ?\' (C X m)" have "M - e \ ?\' (C X m)" by (rule less_imp_le) fact also have "\ \ ?\' (E - (E \ \{k\snd`p. k \ C X m = {}}))" proof (intro outer_measure_of_mono subsetI) fix v assume "v \ C X m" then have "v \ cbox x y" "v \ E" using \E \ space ?L\ \X \ E\ by (auto simp: space_restrict_space C_def) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto then show "v \ E - E \ (\{k\snd`p. k \ C X m = {}})" using \v \ C X m\ \v \ E\ by auto qed also have "\ = ?\ E - ?\ (E \ \{k\snd`p. k \ C X m = {}})" using \E \ sets ?L\ fin[of X] sets[of X] by (auto intro!: emeasure_Diff) finally have "?\ (E \ \{k\snd`p. k \ C X m = {}}) \ e" using \0 < e\ e_less_M apply (cases "?\ (E \ \{k\snd`p. k \ C X m = {}})") by (auto simp add: \?\ E = M\ ennreal_minus ennreal_le_iff2) note this } note upper_bound = this have "?\ (E \ \(snd`(p - s))) = ?\ ((E \ \{k\snd`p. k \ C ?E m = {}}) \ (E \ \{k\snd`p. k \ C ?F m = {}}))" by (intro arg_cong[where f="?\"]) (auto simp: s_def image_def Bex_def) also have "\ \ ?\ (E \ \{k\snd`p. k \ C ?E m = {}}) + ?\ (E \ \{k\snd`p. k \ C ?F m = {}})" using sets[of ?E] sets[of ?F] M_minus_e by (intro emeasure_subadditive) auto also have "\ \ e + ennreal e" using upper_bound[of ?E] upper_bound[of ?F] M_minus_e by (intro add_mono) auto finally have "?\ E - 2*e \ ?\ (E - (E \ \(snd`(p - s))))" using \0 < e\ \E \ sets ?L\ tagged_division_ofD(1)[OF p(1)] by (subst emeasure_Diff) (auto simp: top_unique simp flip: ennreal_plus intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ \ ?\ (\x\p \ s. snd x)" proof (safe intro!: emeasure_mono subsetI) fix v assume "v \ E" and not: "v \ (\x\p \ s. snd x)" then have "v \ cbox x y" using \E \ space ?L\ by (auto simp: space_restrict_space) then obtain z k where "(z, k) \ p" "v \ k" using tagged_division_ofD(6)[OF p(1), symmetric] by auto with not show "v \ \(snd ` (p - s))" by (auto intro!: bexI[of _ "(z, k)"] elim: ballE[of _ _ "(z, k)"]) qed (auto intro!: sets.Int sets.finite_UN ennreal_mono_minus intro: p_in_L) also have "\ = measure ?L (\x\p \ s. snd x)" by (auto intro!: emeasure_eq_ennreal_measure) finally have "M - 2 * e \ measure ?L (\x\p \ s. snd x)" unfolding \?\ E = M\ using \0 < e\ by (simp add: ennreal_minus) also have "measure ?L (\x\p \ s. snd x) = content (\x\p \ s. snd x)" using tagged_division_ofD(1,3,4) [OF p(1)] by (intro content_eq_L[symmetric]) (fastforce intro!: sets.finite_UN UN_least del: subsetI)+ also have "content (\x\p \ s. snd x) \ (\k\p \ s. content (snd k))" using p(1) by (auto simp: emeasure_lborel_cbox_eq intro!: measure_subadditive_finite dest!: p(1)[THEN tagged_division_ofD(4)]) finally show "M - 3 * e \ (\(x, y)\p \ s. content y)" using \0 < e\ by (simp add: split_beta) qed (use \a < b\ in auto) also have "\ = (\(x,k) \ p \ s. content k * (b - a))" by (simp add: sum_distrib_right split_beta') also have "\ \ (\(x,k) \ p \ s. content k * (f (T ?F k) - f (T ?E k)))" using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono) also have "\ = (\(x,k) \ p \ s. content k * f (T ?F k)) - (\(x,k) \ p \ s. content k * f (T ?E k))" by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric]) also have "\ = (\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x)" by (subst (1 2) parts) auto also have "\ \ norm ((\(x,k) \ ?B. content k *\<^sub>R f x) - (\(x,k) \ ?A. content k *\<^sub>R f x))" by auto also have "\ \ e + e" using parts(1)[of ?E] parts(1)[of ?F] by (intro norm_diff_triangle_le[of _ I]) auto finally show False using \2 * e < (b - a) * (M - e * 3)\ by (auto simp: field_simps) qed moreover have "?\' ?E \ ?\ E" "?\' ?F \ ?\ E" unfolding outer_measure_of_eq[OF \E \ sets ?L\, symmetric] by (auto intro!: outer_measure_of_mono) ultimately show "min (?\' ?E) (?\' ?F) < ?\ E" unfolding min_less_iff_disj by (auto simp: less_le) qed lemma has_integral_implies_lebesgue_measurable_real: fixes f :: "'a :: euclidean_space \ real" assumes f: "(f has_integral I) \" shows "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof - define B :: "nat \ 'a set" where "B n = cbox (- real n *\<^sub>R One) (real n *\<^sub>R One)" for n show "(\x. f x * indicator \ x) \ lebesgue \\<^sub>M borel" proof (rule measurable_piecewise_restrict) have "(\n. box (- real n *\<^sub>R One) (real n *\<^sub>R One)) \ \(B ` UNIV)" unfolding B_def by (intro UN_mono box_subset_cbox order_refl) then show "countable (range B)" "space lebesgue \ \(B ` UNIV)" by (auto simp: B_def UN_box_eq_UNIV) next fix \' assume "\' \ range B" then obtain n where \': "\' = B n" by auto then show "\' \ space lebesgue \ sets lebesgue" by (auto simp: B_def) have "f integrable_on \" using f by auto then have "(\x. f x * indicator \ x) integrable_on \" by (auto simp: integrable_on_def cong: has_integral_cong) then have "(\x. f x * indicator \ x) integrable_on (\ \ B n)" by (rule integrable_on_superset) auto then have "(\x. f x * indicator \ x) integrable_on B n" unfolding B_def by (rule integrable_on_subcbox) auto then show "(\x. f x * indicator \ x) \ lebesgue_on \' \\<^sub>M borel" unfolding B_def \' by (auto intro: has_integral_implies_lebesgue_measurable_cbox simp: integrable_on_def) qed qed lemma has_integral_implies_lebesgue_measurable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "(f has_integral I) \" shows "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" proof (intro borel_measurable_euclidean_space[where 'c='b, THEN iffD2] ballI) fix i :: "'b" assume "i \ Basis" have "(\x. (f x \ i) * indicator \ x) \ borel_measurable (completion lborel)" using has_integral_linear[OF f bounded_linear_inner_left, of i] by (intro has_integral_implies_lebesgue_measurable_real) (auto simp: comp_def) then show "(\x. indicator \ x *\<^sub>R f x \ i) \ borel_measurable (completion lborel)" by (simp add: ac_simps) qed subsection \Equivalence Lebesgue integral on \<^const>\lborel\ and HK-integral\ lemma has_integral_measure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and finite: "emeasure lborel A < \" shows "((\x. 1) has_integral measure lborel A) A" proof - { fix l u :: 'a have "((\x. 1) has_integral measure lborel (box l u)) (box l u)" proof cases assume "\b\Basis. l \ b \ u \ b" then show ?thesis apply simp apply (subst has_integral_restrict[symmetric, OF box_subset_cbox]) apply (subst has_integral_spike_interior_eq[where g="\_. 1"]) using has_integral_const[of "1::real" l u] apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases) done next assume "\ (\b\Basis. l \ b \ u \ b)" then have "box l u = {}" unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le) then show ?thesis by simp qed } note has_integral_box = this { fix a b :: 'a let ?M = "\A. measure lborel (A \ box a b)" have "Int_stable (range (\(a, b). box a b))" by (auto simp: Int_stable_def box_Int_box) moreover have "(range (\(a, b). box a b)) \ Pow UNIV" by auto moreover have "A \ sigma_sets UNIV (range (\(a, b). box a b))" using A unfolding borel_eq_box by simp ultimately have "((\x. 1) has_integral ?M A) (A \ box a b)" proof (induction rule: sigma_sets_induct_disjoint) case (basic A) then show ?case by (auto simp: box_Int_box has_integral_box) next case empty then show ?case by simp next case (compl A) then have [measurable]: "A \ sets borel" by (simp add: borel_eq_box) have "((\x. 1) has_integral ?M (box a b)) (box a b)" by (simp add: has_integral_box) moreover have "((\x. if x \ A \ box a b then 1 else 0) has_integral ?M A) (box a b)" by (subst has_integral_restrict) (auto intro: compl) ultimately have "((\x. 1 - (if x \ A \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_diff) then have "((\x. (if x \ (UNIV - A) \ box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)" by (rule has_integral_cong[THEN iffD1, rotated 1]) auto then have "((\x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \ box a b)" by (subst (asm) has_integral_restrict) auto also have "?M (box a b) - ?M A = ?M (UNIV - A)" by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2) finally show ?case . next case (union F) then have [measurable]: "\i. F i \ sets borel" by (simp add: borel_eq_box subset_eq) have "((\x. if x \ \(F ` UNIV) \ box a b then 1 else 0) has_integral ?M (\i. F i)) (box a b)" proof (rule has_integral_monotone_convergence_increasing) let ?f = "\k x. \i F i \ box a b then 1 else 0 :: real" show "\k. (?f k has_integral (\ik x. ?f k x \ ?f (Suc k) x" by (intro sum_mono2) auto from union(1) have *: "\x i j. x \ F i \ x \ F j \ j = i" by (auto simp add: disjoint_family_on_def) show "\x. (\k. ?f k x) \ (if x \ \(F ` UNIV) \ box a b then 1 else 0)" apply (auto simp: * sum.If_cases Iio_Int_singleton) apply (rule_tac k="Suc xa" in LIMSEQ_offset) apply simp done have *: "emeasure lborel ((\x. F x) \ box a b) \ emeasure lborel (box a b)" by (intro emeasure_mono) auto with union(1) show "(\k. \i ?M (\i. F i)" unfolding sums_def[symmetric] UN_extend_simps by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique) qed then show ?case by (subst (asm) has_integral_restrict) auto qed } note * = this show ?thesis proof (rule has_integral_monotone_convergence_increasing) let ?B = "\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set" let ?f = "\n::nat. \x. if x \ A \ ?B n then 1 else 0 :: real" let ?M = "\n. measure lborel (A \ ?B n)" show "\n::nat. (?f n has_integral ?M n) A" using * by (subst has_integral_restrict) simp_all show "\k x. ?f k x \ ?f (Suc k) x" by (auto simp: box_def) { fix x assume "x \ A" moreover have "(\k. indicator (A \ ?B k) x :: real) \ indicator (\k::nat. A \ ?B k) x" by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def) ultimately show "(\k. if x \ A \ ?B k then 1 else 0::real) \ 1" by (simp add: indicator_def UN_box_eq_UNIV) } have "(\n. emeasure lborel (A \ ?B n)) \ emeasure lborel (\n::nat. A \ ?B n)" by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def) also have "(\n. emeasure lborel (A \ ?B n)) = (\n. measure lborel (A \ ?B n))" proof (intro ext emeasure_eq_ennreal_measure) fix n have "emeasure lborel (A \ ?B n) \ emeasure lborel (?B n)" by (intro emeasure_mono) auto then show "emeasure lborel (A \ ?B n) \ top" by (auto simp: top_unique) qed finally show "(\n. measure lborel (A \ ?B n)) \ measure lborel A" using emeasure_eq_ennreal_measure[of lborel A] finite by (simp add: UN_box_eq_UNIV less_top) qed qed lemma nn_integral_has_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" shows "(f has_integral r) UNIV" using f proof (induct f arbitrary: r rule: borel_measurable_induct_real) case (set A) then have "((\x. 1) has_integral measure lborel A) A" by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator) with set show ?case by (simp add: ennreal_indicator measure_def) (simp add: indicator_def) next case (mult g c) then have "ennreal c * (\\<^sup>+ x. g x \lborel) = ennreal r" by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult) with \0 \ r\ \0 \ c\ obtain r' where "(c = 0 \ r = 0) \ (0 \ r' \ (\\<^sup>+ x. ennreal (g x) \lborel) = ennreal r' \ r = c * r')" by (cases "\\<^sup>+ x. ennreal (g x) \lborel" rule: ennreal_cases) (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric]) with mult show ?case by (auto intro!: has_integral_cmult_real) next case (add g h) then have "(\\<^sup>+ x. h x + g x \lborel) = (\\<^sup>+ x. h x \lborel) + (\\<^sup>+ x. g x \lborel)" by (simp add: nn_integral_add) with add obtain a b where "0 \ a" "0 \ b" "(\\<^sup>+ x. h x \lborel) = ennreal a" "(\\<^sup>+ x. g x \lborel) = ennreal b" "r = a + b" by (cases "\\<^sup>+ x. h x \lborel" "\\<^sup>+ x. g x \lborel" rule: ennreal2_cases) (auto simp: add_top nn_integral_add top_add simp flip: ennreal_plus) with add show ?case by (auto intro!: has_integral_add) next case (seq U) note seq(1)[measurable] and f[measurable] { fix i x have "U i x \ f x" using seq(5) apply (rule LIMSEQ_le_const) using seq(4) apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def) done } note U_le_f = this { fix i have "(\\<^sup>+x. U i x \lborel) \ (\\<^sup>+x. f x \lborel)" using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp then obtain p where "(\\<^sup>+x. U i x \lborel) = ennreal p" "p \ r" "0 \ p" using seq(6) \0\r\ by (cases "\\<^sup>+x. U i x \lborel" rule: ennreal_cases) (auto simp: top_unique) moreover note seq ultimately have "\p. (\\<^sup>+x. U i x \lborel) = ennreal p \ 0 \ p \ p \ r \ (U i has_integral p) UNIV" by auto } then obtain p where p: "\i. (\\<^sup>+x. ennreal (U i x) \lborel) = ennreal (p i)" and bnd: "\i. p i \ r" "\i. 0 \ p i" and U_int: "\i.(U i has_integral (p i)) UNIV" by metis have int_eq: "\i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) \ integral UNIV f" proof (rule monotone_convergence_increasing) show "\k. U k integrable_on UNIV" using U_int by auto show "\k x. x\UNIV \ U k x \ U (Suc k) x" using \incseq U\ by (auto simp: incseq_def le_fun_def) then show "bounded (range (\k. integral UNIV (U k)))" using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) show "\x. x\UNIV \ (\k. U k x) \ f x" using seq by auto qed moreover have "(\i. (\\<^sup>+x. U i x \lborel)) \ (\\<^sup>+x. f x \lborel)" using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto ultimately have "integral UNIV f = r" by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique) with * show ?case by (simp add: has_integral_integral) qed lemma nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "(\\<^sup>+x. f x \lborel) = integral UNIV f" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique) qed lemma nn_integral_integrable_on: fixes f::"'a::euclidean_space \ real" assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" "(\\<^sup>+x. f x \lborel) < \" shows "f integrable_on UNIV" proof - from f(3) obtain r where r: "(\\<^sup>+x. f x \lborel) = ennreal r" "0 \ r" by (cases "\\<^sup>+x. f x \lborel" rule: ennreal_cases) auto then show ?thesis by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f) qed lemma nn_integral_has_integral_lborel: fixes f :: "'a::euclidean_space \ real" assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" shows "integral\<^sup>N lborel f = I" proof - from f_borel have "(\x. ennreal (f x)) \ borel_measurable lborel" by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain F where F: "\i. simple_function lborel (F i)" "incseq F" "\i x. F i x < top" "\x. (SUP i. F i x) = ennreal (f x)" by blast then have [measurable]: "\i. F i \ borel_measurable lborel" by (metis borel_measurable_simple_function) let ?B = "\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set" have "0 \ I" using I by (rule has_integral_nonneg) (simp add: nonneg) have F_le_f: "enn2real (F i x) \ f x" for i x using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\i. F i x"] by (cases "F i x" rule: ennreal_cases) auto let ?F = "\i x. F i x * indicator (?B i) x" have "(\\<^sup>+ x. ennreal (f x) \lborel) = (SUP i. integral\<^sup>N lborel (\x. ?F i x))" proof (subst nn_integral_monotone_convergence_SUP[symmetric]) { fix x obtain j where j: "x \ ?B j" using UN_box_eq_UNIV by auto have "ennreal (f x) = (SUP i. F i x)" using F(4)[of x] nonneg[of x] by (simp add: max_def) also have "\ = (SUP i. ?F i x)" proof (rule SUP_eq) fix i show "\j\UNIV. F i x \ ?F j x" using j F(2) by (intro bexI[of _ "max i j"]) (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def) qed (auto intro!: F split: split_indicator) finally have "ennreal (f x) = (SUP i. ?F i x)" . } then show "(\\<^sup>+ x. ennreal (f x) \lborel) = (\\<^sup>+ x. (SUP i. ?F i x) \lborel)" by simp qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator) also have "\ \ ennreal I" proof (rule SUP_least) fix i :: nat have finite_F: "(\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel) < \" proof (rule nn_integral_bound_simple_function) have "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} \ emeasure lborel (?B i)" by (intro emeasure_mono) (auto split: split_indicator) then show "emeasure lborel {x \ space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \ 0} < \" by (auto simp: less_top[symmetric] top_unique) qed (auto split: split_indicator intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal) have int_F: "(\x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV" using F(4) finite_F by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg) have "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) = (\\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \lborel)" using F(3,4) by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator) also have "\ = ennreal (integral UNIV (\x. enn2real (F i x) * indicator (?B i) x))" using F by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F]) (auto split: split_indicator intro: enn2real_nonneg) also have "\ \ ennreal I" by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f simp: \0 \ I\ split: split_indicator ) finally show "(\\<^sup>+ x. F i x * indicator (?B i) x \lborel) \ ennreal I" . qed finally have "(\\<^sup>+ x. ennreal (f x) \lborel) < \" by (auto simp: less_top[symmetric] top_unique) from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis by (simp add: integral_unique) qed lemma has_integral_iff_emeasure_lborel: fixes A :: "'a::euclidean_space set" assumes A[measurable]: "A \ sets borel" and [simp]: "0 \ r" shows "((\x. 1) has_integral r) A \ emeasure lborel A = ennreal r" proof (cases "emeasure lborel A = \") case emeasure_A: True have "\ (\x. 1::real) integrable_on A" proof assume int: "(\x. 1::real) integrable_on A" then have "(indicator A::'a \ real) integrable_on UNIV" unfolding indicator_def[abs_def] integrable_restrict_UNIV . then obtain r where "((indicator A::'a\real) has_integral r) UNIV" by auto from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False by (simp add: ennreal_indicator) qed with emeasure_A show ?thesis by auto next case False then have "((\x. 1) has_integral measure lborel A) A" by (simp add: has_integral_measure_lborel less_top) with False show ?thesis by (auto simp: emeasure_eq_ennreal_measure has_integral_unique) qed lemma ennreal_max_0: "ennreal (max 0 x) = ennreal x" by (auto simp: max_def ennreal_neg) lemma has_integral_integral_real: fixes f::"'a::euclidean_space \ real" assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - from integrableE[OF f] obtain r q where "0 \ r" "0 \ q" and r: "(\\<^sup>+ x. ennreal (max 0 (f x)) \lborel) = ennreal r" and q: "(\\<^sup>+ x. ennreal (max 0 (- f x)) \lborel) = ennreal q" and f: "f \ borel_measurable lborel" and eq: "integral\<^sup>L lborel f = r - q" unfolding ennreal_max_0 by auto then have "((\x. max 0 (f x)) has_integral r) UNIV" "((\x. max 0 (- f x)) has_integral q) UNIV" using nn_integral_has_integral[OF _ _ r] nn_integral_has_integral[OF _ _ q] by auto note has_integral_diff[OF this] moreover have "(\x. max 0 (f x) - max 0 (- f x)) = f" by auto ultimately show ?thesis by (simp add: eq) qed lemma has_integral_AE: assumes ae: "AE x in lborel. x \ \ \ f x = g x" shows "(f has_integral x) \ = (g has_integral x) \" proof - from ae obtain N where N: "N \ sets borel" "emeasure lborel N = 0" "{x. \ (x \ \ \ f x = g x)} \ N" by (auto elim!: AE_E) then have not_N: "AE x in lborel. x \ N" by (simp add: AE_iff_measurable) show ?thesis proof (rule has_integral_spike_eq[symmetric]) show "\x. x\\ - N \ f x = g x" using N(3) by auto show "negligible N" unfolding negligible_def proof (intro allI) fix a b :: "'a" let ?F = "\x::'a. if x \ cbox a b then indicator N x else 0 :: real" have "integrable lborel ?F = integrable lborel (\x::'a. 0::real)" using not_N N(1) by (intro integrable_cong_AE) auto moreover have "(LINT x|lborel. ?F x) = (LINT x::'a|lborel. 0::real)" using not_N N(1) by (intro integral_cong_AE) auto ultimately have "(?F has_integral 0) UNIV" using has_integral_integral_real[of ?F] by simp then show "(indicator N has_integral (0::real)) (cbox a b)" unfolding has_integral_restrict_UNIV . qed qed qed lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto finally show ?thesis using nonneg by auto qed lemma has_integral_iff_nn_integral_lebesgue: assumes f: "\x. 0 \ f x" shows "(f has_integral r) UNIV \ (f \ lebesgue \\<^sub>M borel \ integral\<^sup>N lebesgue f = r \ 0 \ r)" (is "?I = ?N") proof assume ?I have "0 \ r" using has_integral_nonneg[OF \?I\] f by auto then show ?N using nn_integral_has_integral_lebesgue[OF f \?I\] has_integral_implies_lebesgue_measurable[OF \?I\] by (auto simp: nn_integral_completion) next assume ?N then obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" by (auto dest: completion_ex_borel_measurable_real) moreover have "(\\<^sup>+ x. ennreal \f' x\ \lborel) = (\\<^sup>+ x. ennreal \f x\ \lborel)" using f' by (intro nn_integral_cong_AE) auto moreover have "((\x. \f' x\) has_integral r) UNIV \ ((\x. \f x\) has_integral r) UNIV" using f' by (intro has_integral_AE) auto moreover note nn_integral_has_integral[of "\x. \f' x\" r] \?N\ ultimately show ?I using f by (auto simp: nn_integral_completion) qed context fixes f::"'a::euclidean_space \ 'b::euclidean_space" begin lemma has_integral_integral_lborel: assumes f: "integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lborel (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lborel f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma integrable_on_lborel: "integrable lborel f \ f integrable_on UNIV" using has_integral_integral_lborel by auto lemma integral_lborel: "integrable lborel f \ integral UNIV f = (\x. f x \lborel)" using has_integral_integral_lborel by auto end context begin private lemma has_integral_integral_lebesgue_real: fixes f :: "'a::euclidean_space \ real" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - obtain f' where f': "f' \ borel \\<^sub>M borel" "AE x in lborel. f x = f' x" using completion_ex_borel_measurable_real[OF borel_measurable_integrable[OF f]] by auto moreover have "(\\<^sup>+ x. ennreal (norm (f x)) \lborel) = (\\<^sup>+ x. ennreal (norm (f' x)) \lborel)" using f' by (intro nn_integral_cong_AE) auto ultimately have "integrable lborel f'" using f by (auto simp: integrable_iff_bounded nn_integral_completion cong: nn_integral_cong_AE) note has_integral_integral_real[OF this] moreover have "integral\<^sup>L lebesgue f = integral\<^sup>L lebesgue f'" using f' f by (intro integral_cong_AE) (auto intro: AE_completion measurable_completion) moreover have "integral\<^sup>L lebesgue f' = integral\<^sup>L lborel f'" using f' by (simp add: integral_completion) moreover have "(f' has_integral integral\<^sup>L lborel f') UNIV \ (f has_integral integral\<^sup>L lborel f') UNIV" using f' by (intro has_integral_AE) auto ultimately show ?thesis by auto qed lemma has_integral_integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" proof - have "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b)) UNIV" using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto also have eq_f: "(\x. \b\Basis. (f x \ b) *\<^sub>R b) = f" by (simp add: fun_eq_iff euclidean_representation) also have "(\b\Basis. integral\<^sup>L lebesgue (\x. f x \ b) *\<^sub>R b) = integral\<^sup>L lebesgue f" using f by (subst (2) eq_f[symmetric]) simp finally show ?thesis . qed lemma has_integral_integral_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" proof - let ?f = "\x. if x \ S then f x else 0" have "integrable lebesgue (\x. indicat_real S x *\<^sub>R f x)" using indicator_scaleR_eq_if [of S _ f] assms by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) then have "integrable lebesgue ?f" using indicator_scaleR_eq_if [of S _ f] assms by auto then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" by (rule has_integral_integral_lebesgue) then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" using has_integral_restrict_UNIV by blast moreover have "S \ space lebesgue \ sets lebesgue" by (simp add: assms) then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" by (simp add: integral_restrict_space indicator_scaleR_eq_if) ultimately show ?thesis by auto qed lemma lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = integral S f" by (metis has_integral_integral_lebesgue_on assms integral_unique) lemma integrable_on_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ f integrable_on UNIV" using has_integral_integral_lebesgue by auto lemma integral_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "integrable lebesgue f \ integral UNIV f = (\x. f x \lebesgue)" using has_integral_integral_lebesgue by auto end subsection \Absolute integrability (this is the same as Lebesgue integrability)\ translations "LBINT x. f" == "CONST lebesgue_integral CONST lborel (\x. f)" translations "LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\x. f)" lemma set_integral_reflect: fixes S and f :: "real \ 'a :: {banach, second_countable_topology}" shows "(LBINT x : S. f x) = (LBINT x : {x. - x \ S}. f (- x))" unfolding set_lebesgue_integral_def by (subst lborel_integral_real_affine[where c="-1" and t=0]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma borel_integrable_atLeastAtMost': fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes f: "continuous_on {a..b} f" shows "set_integrable lborel {a..b} f" unfolding set_integrable_def by (intro borel_integrable_compact compact_Icc f) lemma integral_FTC_atLeastAtMost: fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" and F: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" and f: "continuous_on {a .. b} f" shows "integral\<^sup>L lborel (\x. indicator {a .. b} x *\<^sub>R f x) = F b - F a" proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have "(?f has_integral (\x. ?f x \lborel)) UNIV" using borel_integrable_atLeastAtMost'[OF f] unfolding set_integrable_def by (rule has_integral_integral_lborel) moreover have "(f has_integral F b - F a) {a .. b}" by (intro fundamental_theorem_of_calculus ballI assms) auto then have "(?f has_integral F b - F a) {a .. b}" by (subst has_integral_cong[where g=f]) auto then have "(?f has_integral F b - F a) UNIV" by (intro has_integral_on_superset[where T=UNIV and S="{a..b}"]) auto ultimately show "integral\<^sup>L lborel ?f = F b - F a" by (rule has_integral_unique) qed lemma set_borel_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "set_integrable lborel S f" shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f" proof - let ?f = "\x. indicator S x *\<^sub>R f x" have "(?f has_integral LINT x : S | lborel. f x) UNIV" using assms has_integral_integral_lborel unfolding set_integrable_def set_lebesgue_integral_def by blast hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S" apply (subst has_integral_restrict_UNIV [symmetric]) apply (rule has_integral_eq) by auto thus "f integrable_on S" by (auto simp add: integrable_on_def) with 1 have "(f has_integral (integral S f)) S" by (intro integrable_integral, auto simp add: integrable_on_def) thus "LINT x : S | lborel. f x = integral S f" by (intro has_integral_unique [OF 1]) qed lemma has_integral_set_lebesgue: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "(f has_integral (LINT x:S|lebesgue. f x)) S" using has_integral_integral_lebesgue f by (fastforce simp add: set_integrable_def set_lebesgue_integral_def indicator_def if_distrib[of "\x. x *\<^sub>R f _"] cong: if_cong) lemma set_lebesgue_integral_eq_integral: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "set_integrable lebesgue S f" shows "f integrable_on S" "LINT x:S | lebesgue. f x = integral S f" using has_integral_set_lebesgue[OF f] by (auto simp: integral_unique integrable_on_def) lemma lmeasurable_iff_has_integral: "S \ lmeasurable \ ((indicator S) has_integral measure lebesgue S) UNIV" by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: ennreal_indicator emeasure_eq_measure2 borel_measurable_indicator_iff intro!: fmeasurableI) abbreviation absolutely_integrable_on :: "('a::euclidean_space \ 'b::{banach, second_countable_topology}) \ 'a set \ bool" (infixr "absolutely'_integrable'_on" 46) where "f absolutely_integrable_on s \ set_integrable lebesgue s f" lemma absolutely_integrable_zero [simp]: "(\x. 0) absolutely_integrable_on S" by (simp add: set_integrable_def) lemma absolutely_integrable_on_def: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ f integrable_on S \ (\x. norm (f x)) integrable_on S" proof safe assume f: "f absolutely_integrable_on S" then have nf: "integrable lebesgue (\x. norm (indicator S x *\<^sub>R f x))" using integrable_norm set_integrable_def by blast show "f integrable_on S" by (rule set_lebesgue_integral_eq_integral[OF f]) have "(\x. norm (indicator S x *\<^sub>R f x)) = (\x. if x \ S then norm (f x) else 0)" by auto with integrable_on_lebesgue[OF nf] show "(\x. norm (f x)) integrable_on S" by (simp add: integrable_restrict_UNIV) next assume f: "f integrable_on S" and nf: "(\x. norm (f x)) integrable_on S" show "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule integrableI_bounded) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lebesgue" using f has_integral_implies_lebesgue_measurable[of f _ S] by (auto simp: integrable_on_def) show "(\\<^sup>+ x. ennreal (norm (indicator S x *\<^sub>R f x)) \lebesgue) < \" using nf nn_integral_has_integral_lebesgue[of "\x. norm (f x)" _ S] by (auto simp: integrable_on_def nn_integral_completion) qed qed lemma integrable_on_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "integrable (lebesgue_on S) f" and S: "S \ sets lebesgue" shows "f integrable_on S" proof - have "integrable lebesgue (\x. indicator S x *\<^sub>R f x)" using S f inf_top.comm_neutral integrable_restrict_space by blast then show ?thesis using absolutely_integrable_on_def set_integrable_def by blast qed lemma absolutely_integrable_imp_integrable: assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) lemma absolutely_integrable_on_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "content (cbox a b) = 0 \ f absolutely_integrable_on (cbox a b)" by (auto simp: absolutely_integrable_on_def) lemma absolutely_integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f absolutely_integrable_on box a b \ f absolutely_integrable_on cbox a b" by (auto simp: integrable_on_open_interval absolutely_integrable_on_def) lemma absolutely_integrable_restrict_UNIV: "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ f absolutely_integrable_on S" unfolding set_integrable_def by (intro arg_cong2[where f=integrable]) auto lemma absolutely_integrable_onI: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f integrable_on S \ (\x. norm (f x)) integrable_on S \ f absolutely_integrable_on S" unfolding absolutely_integrable_on_def by auto lemma nonnegative_absolutely_integrable_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "\x. x \ A \ 0 \ f x" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_onI [OF f]) (use assms in \simp add: integrable_eq\) lemma absolutely_integrable_on_iff_nonneg: fixes f :: "'a :: euclidean_space \ real" assumes "\x. x \ S \ 0 \ f x" shows "f absolutely_integrable_on S \ f integrable_on S" proof - { assume "f integrable_on S" then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" using \f integrable_on S\ absolutely_integrable_restrict_UNIV assms nonnegative_absolutely_integrable_1 by blast then have "f absolutely_integrable_on S" using absolutely_integrable_restrict_UNIV by blast } then show ?thesis unfolding absolutely_integrable_on_def by auto qed lemma absolutely_integrable_on_scaleR_iff: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S \ c = 0 \ f absolutely_integrable_on S" proof (cases "c=0") case False then show ?thesis unfolding absolutely_integrable_on_def by (simp add: norm_mult) qed auto lemma absolutely_integrable_spike: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on T" and S: "negligible S" "\x. x \ T - S \ g x = f x" shows "g absolutely_integrable_on T" using assms integrable_spike unfolding absolutely_integrable_on_def by metis lemma absolutely_integrable_negligible: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" shows "f absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def integrable_negligible) lemma absolutely_integrable_spike_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible S" "\x. x \ T - S \ g x = f x" shows "(f absolutely_integrable_on T \ g absolutely_integrable_on T)" using assms by (blast intro: absolutely_integrable_spike sym) lemma absolutely_integrable_spike_set_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f absolutely_integrable_on S \ f absolutely_integrable_on T)" proof - have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV \ (\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" proof (rule absolutely_integrable_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto with absolutely_integrable_restrict_UNIV show ?thesis by blast qed lemma absolutely_integrable_spike_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f absolutely_integrable_on T" using absolutely_integrable_spike_set_eq f neg by blast lemma absolutely_integrable_reflect[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \ f absolutely_integrable_on cbox a b" apply (auto simp: absolutely_integrable_on_def dest: integrable_reflect [THEN iffD1]) unfolding integrable_on_def by auto lemma absolutely_integrable_reflect_real[simp]: fixes f :: "real \ 'b::euclidean_space" shows "(\x. f(-x)) absolutely_integrable_on {-b .. -a} \ f absolutely_integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) lemma absolutely_integrable_on_subcbox: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; cbox a b \ S\ \ f absolutely_integrable_on cbox a b" by (meson absolutely_integrable_on_def integrable_on_subcbox) lemma absolutely_integrable_on_subinterval: fixes f :: "real \ 'b::euclidean_space" shows "\f absolutely_integrable_on S; {a..b} \ S\ \ f absolutely_integrable_on {a..b}" using absolutely_integrable_on_subcbox by fastforce lemma integrable_subinterval: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" and "{c..d} \ {a..b}" shows "integrable (lebesgue_on {c..d}) f" proof (rule absolutely_integrable_imp_integrable) show "f absolutely_integrable_on {c..d}" proof - have "f integrable_on {c..d}" using assms integrable_on_lebesgue_on integrable_on_subinterval by fastforce moreover have "(\x. norm (f x)) integrable_on {c..d}" proof (rule integrable_on_subinterval) show "(\x. norm (f x)) integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) qed (use assms in auto) ultimately show ?thesis by (auto simp: absolutely_integrable_on_def) qed qed auto lemma indefinite_integral_continuous_real: fixes f :: "real \ 'a::euclidean_space" assumes "integrable (lebesgue_on {a..b}) f" shows "continuous_on {a..b} (\x. integral\<^sup>L (lebesgue_on {a..x}) f)" proof - have "f integrable_on {a..b}" by (simp add: assms integrable_on_lebesgue_on) then have "continuous_on {a..b} (\x. integral {a..x} f)" using indefinite_integral_continuous_1 by blast moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \ x" "x \ b" for x proof - have "{a..x} \ {a..b}" using that by auto then have "integrable (lebesgue_on {a..x}) f" using integrable_subinterval assms by blast then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" by (simp add: lebesgue_integral_eq_integral) qed ultimately show ?thesis by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) qed lemma lmeasurable_iff_integrable_on: "S \ lmeasurable \ (\x. 1::real) integrable_on S" by (subst absolutely_integrable_on_iff_nonneg[symmetric]) (simp_all add: lmeasurable_iff_integrable set_integrable_def) lemma lmeasure_integral_UNIV: "S \ lmeasurable \ measure lebesgue S = integral UNIV (indicator S)" by (simp add: lmeasurable_iff_has_integral integral_unique) lemma lmeasure_integral: "S \ lmeasurable \ measure lebesgue S = integral S (\x. 1::real)" by (fastforce simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) lemma integrable_on_const: "S \ lmeasurable \ (\x. c) integrable_on S" unfolding lmeasurable_iff_integrable by (metis (mono_tags, lifting) integrable_eq integrable_on_scaleR_left lmeasurable_iff_integrable lmeasurable_iff_integrable_on scaleR_one) lemma integral_indicator: assumes "(S \ T) \ lmeasurable" shows "integral T (indicator S) = measure lebesgue (S \ T)" proof - have "integral UNIV (indicator (S \ T)) = integral UNIV (\a. if a \ S \ T then 1::real else 0)" by (meson indicator_def) moreover have "(indicator (S \ T) has_integral measure lebesgue (S \ T)) UNIV" using assms by (simp add: lmeasurable_iff_has_integral) ultimately have "integral UNIV (\x. if x \ S \ T then 1 else 0) = measure lebesgue (S \ T)" by (metis (no_types) integral_unique) then show ?thesis using integral_restrict_Int [of UNIV "S \ T" "\x. 1::real"] apply (simp add: integral_restrict_Int [symmetric]) by (meson indicator_def) qed lemma measurable_integrable: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ (indicat_real S) integrable_on UNIV" by (auto simp: lmeasurable_iff_integrable absolutely_integrable_on_iff_nonneg [symmetric] set_integrable_def) lemma integrable_on_indicator: fixes S :: "'a::euclidean_space set" shows "indicat_real S integrable_on T \ (S \ T) \ lmeasurable" unfolding measurable_integrable unfolding integrable_restrict_UNIV [of T, symmetric] by (fastforce simp add: indicator_def elim: integrable_eq) lemma assumes \: "\ division_of S" shows lmeasurable_division: "S \ lmeasurable" (is ?l) and content_division: "(\k\\. measure lebesgue k) = measure lebesgue S" (is ?m) proof - { fix d1 d2 assume *: "d1 \ \" "d2 \ \" "d1 \ d2" then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" using division_ofD(4)[OF \] by blast with division_ofD(5)[OF \ *] have "d1 \ sets lborel" "d2 \ sets lborel" "d1 \ d2 \ (cbox a b - box a b) \ (cbox c d - box c d)" by auto moreover have "(cbox a b - box a b) \ (cbox c d - box c d) \ null_sets lborel" by (intro null_sets.Un null_sets_cbox_Diff_box) ultimately have "d1 \ d2 \ null_sets lborel" by (blast intro: null_sets_subset) } then show ?l ?m unfolding division_ofD(6)[OF \, symmetric] using division_ofD(1,4)[OF \] by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) qed lemma has_measure_limit: assumes "S \ lmeasurable" "e > 0" obtains B where "B > 0" "\a b. ball 0 B \ cbox a b \ \measure lebesgue (S \ cbox a b) - measure lebesgue S\ < e" using assms unfolding lmeasurable_iff_has_integral has_integral_alt' by (force simp: integral_indicator integrable_on_indicator) lemma lmeasurable_iff_indicator_has_integral: fixes S :: "'a::euclidean_space set" shows "S \ lmeasurable \ m = measure lebesgue S \ (indicat_real S has_integral m) UNIV" by (metis has_integral_iff lmeasurable_iff_has_integral measurable_integrable) lemma has_measure_limit_iff: fixes f :: "'n::euclidean_space \ 'a::banach" shows "S \ lmeasurable \ m = measure lebesgue S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (S \ cbox a b) \ lmeasurable \ \measure lebesgue (S \ cbox a b) - m\ < e)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson has_measure_limit fmeasurable.Int lmeasurable_cbox) next assume RHS [rule_format]: ?rhs show ?lhs apply (simp add: lmeasurable_iff_indicator_has_integral has_integral' [where i=m]) using RHS by (metis (full_types) integral_indicator integrable_integral integrable_on_indicator) qed subsection\Applications to Negligibility\ lemma negligible_iff_null_sets: "negligible S \ S \ null_sets lebesgue" proof assume "negligible S" then have "(indicator S has_integral (0::real)) UNIV" by (auto simp: negligible) then show "S \ null_sets lebesgue" by (subst (asm) has_integral_iff_nn_integral_lebesgue) (auto simp: borel_measurable_indicator_iff nn_integral_0_iff_AE AE_iff_null_sets indicator_eq_0_iff) next assume S: "S \ null_sets lebesgue" show "negligible S" unfolding negligible_def proof (safe intro!: has_integral_iff_nn_integral_lebesgue[THEN iffD2] has_integral_restrict_UNIV[where s="cbox _ _", THEN iffD1]) fix a b show "(\x. if x \ cbox a b then indicator S x else 0) \ lebesgue \\<^sub>M borel" using S by (auto intro!: measurable_If) then show "(\\<^sup>+ x. ennreal (if x \ cbox a b then indicator S x else 0) \lebesgue) = ennreal 0" using S[THEN AE_not_in] by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) qed auto qed corollary eventually_ae_filter_negligible: "eventually P (ae_filter lebesgue) \ (\N. negligible N \ {x. \ P x} \ N)" by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) lemma starlike_negligible: assumes "closed S" and eq1: "\c x. (a + c *\<^sub>R x) \ S \ 0 \ c \ a + x \ S \ c = 1" shows "negligible S" proof - have "negligible ((+) (-a) ` S)" proof (subst negligible_on_intervals, intro allI) fix u v show "negligible ((+) (- a) ` S \ cbox u v)" using \closed S\ eq1 by (auto simp add: negligible_iff_null_sets algebra_simps intro!: closed_translation_subtract starlike_negligible_compact cong: image_cong_simp) (metis add_diff_eq diff_add_cancel scale_right_diff_distrib) qed then show ?thesis by (rule negligible_translation_rev) qed lemma starlike_negligible_strong: assumes "closed S" and star: "\c x. \0 \ c; c < 1; a+x \ S\ \ a + c *\<^sub>R x \ S" shows "negligible S" proof - show ?thesis proof (rule starlike_negligible [OF \closed S\, of a]) fix c x assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" with star have "\ (c < 1)" by auto moreover have "\ (c > 1)" using star [of "1/c" "c *\<^sub>R x"] cx by force ultimately show "c = 1" by arith qed qed lemma negligible_hyperplane: assumes "a \ 0 \ b \ 0" shows "negligible {x. a \ x = b}" proof - obtain x where x: "a \ x \ b" using assms apply auto apply (metis inner_eq_zero_iff inner_zero_right) using inner_zero_right by fastforce show ?thesis apply (rule starlike_negligible [OF closed_hyperplane, of x]) using x apply (auto simp: algebra_simps) done qed lemma negligible_lowdim: fixes S :: "'N :: euclidean_space set" assumes "dim S < DIM('N)" shows "negligible S" proof - obtain a where "a \ 0" and a: "span S \ {x. a \ x = 0}" using lowdim_subset_hyperplane [OF assms] by blast have "negligible (span S)" using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) then show ?thesis using span_base by (blast intro: negligible_subset) qed proposition negligible_convex_frontier: fixes S :: "'N :: euclidean_space set" assumes "convex S" shows "negligible(frontier S)" proof - have nf: "negligible(frontier S)" if "convex S" "0 \ S" for S :: "'N set" proof - obtain B where "B \ S" and indB: "independent B" and spanB: "S \ span B" and cardB: "card B = dim S" by (metis basis_exists) consider "dim S < DIM('N)" | "dim S = DIM('N)" using dim_subset_UNIV le_eq_less_or_eq by auto then show ?thesis proof cases case 1 show ?thesis by (rule negligible_subset [of "closure S"]) (simp_all add: frontier_def negligible_lowdim 1) next case 2 obtain a where a: "a \ interior S" apply (rule interior_simplex_nonempty [OF indB]) apply (simp add: indB independent_finite) apply (simp add: cardB 2) apply (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) done show ?thesis proof (rule starlike_negligible_strong [where a=a]) fix c::real and x have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" by (simp add: algebra_simps) assume "0 \ c" "c < 1" "a + x \ frontier S" then show "a + c *\<^sub>R x \ frontier S" apply (clarsimp simp: frontier_def) apply (subst eq) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"], auto) done qed auto qed qed show ?thesis proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain a where "a \ S" by auto show ?thesis using nf [of "(\x. -a + x) ` S"] by (metis \a \ S\ add.left_inverse assms convex_translation_eq frontier_translation image_eqI negligible_translation_rev) qed qed corollary negligible_sphere: "negligible (sphere a e)" using frontier_cball negligible_convex_frontier convex_cball by (blast intro: negligible_subset) lemma non_negligible_UNIV [simp]: "\ negligible UNIV" unfolding negligible_iff_null_sets by (auto simp: null_sets_def) lemma negligible_interval: "negligible (cbox a b) \ box a b = {}" "negligible (box a b) \ box a b = {}" by (auto simp: negligible_iff_null_sets null_sets_def prod_nonneg inner_diff_left box_eq_empty not_le emeasure_lborel_cbox_eq emeasure_lborel_box_eq intro: eq_refl antisym less_imp_le) proposition open_not_negligible: assumes "open S" "S \ {}" shows "\ negligible S" proof assume neg: "negligible S" obtain a where "a \ S" using \S \ {}\ by blast then obtain e where "e > 0" "cball a e \ S" using \open S\ open_contains_cball_eq by blast let ?p = "a - (e / DIM('a)) *\<^sub>R One" let ?q = "a + (e / DIM('a)) *\<^sub>R One" have "cbox ?p ?q \ cball a e" proof (clarsimp simp: mem_box dist_norm) fix x assume "\i\Basis. ?p \ i \ x \ i \ x \ i \ ?q \ i" then have ax: "\(a - x) \ i\ \ e / real DIM('a)" if "i \ Basis" for i using that by (auto simp: algebra_simps) have "norm (a - x) \ (\i\Basis. \(a - x) \ i\)" by (rule norm_le_l1) also have "\ \ DIM('a) * (e / real DIM('a))" by (intro sum_bounded_above ax) also have "\ = e" by simp finally show "norm (a - x) \ e" . qed then have "negligible (cbox ?p ?q)" by (meson \cball a e \ S\ neg negligible_subset) with \e > 0\ show False by (simp add: negligible_interval box_eq_empty algebra_simps field_split_simps mult_le_0_iff) qed lemma negligible_convex_interior: "convex S \ (negligible S \ interior S = {})" apply safe apply (metis interior_subset negligible_subset open_interior open_not_negligible) apply (metis Diff_empty closure_subset frontier_def negligible_convex_frontier negligible_subset) done lemma measure_eq_0_null_sets: "S \ null_sets M \ measure M S = 0" by (auto simp: measure_def null_sets_def) lemma negligible_imp_measure0: "negligible S \ measure lebesgue S = 0" by (simp add: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_emeasure0: "S \ sets lebesgue \ (negligible S \ emeasure lebesgue S = 0)" by (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) lemma negligible_iff_measure0: "S \ lmeasurable \ (negligible S \ measure lebesgue S = 0)" apply (auto simp: measure_eq_0_null_sets negligible_iff_null_sets) by (metis completion.null_sets_outer subsetI) lemma negligible_imp_sets: "negligible S \ S \ sets lebesgue" by (simp add: negligible_iff_null_sets null_setsD2) lemma negligible_imp_measurable: "negligible S \ S \ lmeasurable" by (simp add: fmeasurableI_null_sets negligible_iff_null_sets) lemma negligible_iff_measure: "negligible S \ S \ lmeasurable \ measure lebesgue S = 0" by (fastforce simp: negligible_iff_measure0 negligible_imp_measurable dest: negligible_imp_measure0) lemma negligible_outer: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T < e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis negligible_iff_measure order_refl) next assume ?rhs then show "negligible S" by (meson completion.null_sets_outer negligible_iff_null_sets) qed lemma negligible_outer_le: "negligible S \ (\e>0. \T. S \ T \ T \ lmeasurable \ measure lebesgue T \ e)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs by (metis dual_order.strict_implies_order negligible_imp_measurable negligible_imp_measure0 order_refl) next assume ?rhs then show "negligible S" by (metis le_less_trans negligible_outer field_lbound_gt_zero) qed lemma negligible_UNIV: "negligible S \ (indicat_real S has_integral 0) UNIV" (is "_=?rhs") proof assume ?rhs then show "negligible S" apply (auto simp: negligible_def has_integral_iff integrable_on_indicator) by (metis negligible integral_unique lmeasure_integral_UNIV negligible_iff_measure0) qed (simp add: negligible) lemma sets_negligible_symdiff: "\S \ sets lebesgue; negligible((S - T) \ (T - S))\ \ T \ sets lebesgue" by (metis Diff_Diff_Int Int_Diff_Un inf_commute negligible_Un_eq negligible_imp_sets sets.Diff sets.Un) lemma lmeasurable_negligible_symdiff: "\S \ lmeasurable; negligible((S - T) \ (T - S))\ \ T \ lmeasurable" using integrable_spike_set_eq lmeasurable_iff_integrable_on by blast lemma measure_Un3_negligible: assumes meas: "S \ lmeasurable" "T \ lmeasurable" "U \ lmeasurable" and neg: "negligible(S \ T)" "negligible(S \ U)" "negligible(T \ U)" and V: "S \ T \ U = V" shows "measure lebesgue V = measure lebesgue S + measure lebesgue T + measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ T) = 0" using neg(1) negligible_imp_measure0 by blast have [simp]: "measure lebesgue (S \ U \ T \ U) = 0" using neg(2) neg(3) negligible_Un negligible_imp_measure0 by blast have "measure lebesgue V = measure lebesgue (S \ T \ U)" using V by simp also have "\ = measure lebesgue S + measure lebesgue T + measure lebesgue U" by (simp add: measure_Un3 meas fmeasurable.Un Int_Un_distrib2) finally show ?thesis . qed lemma measure_translate_add: assumes meas: "S \ lmeasurable" "T \ lmeasurable" and U: "S \ ((+)a ` T) = U" and neg: "negligible(S \ ((+)a ` T))" shows "measure lebesgue S + measure lebesgue T = measure lebesgue U" proof - have [simp]: "measure lebesgue (S \ (+) a ` T) = 0" using neg negligible_imp_measure0 by blast have "measure lebesgue (S \ ((+)a ` T)) = measure lebesgue S + measure lebesgue T" by (simp add: measure_Un3 meas measurable_translation measure_translation fmeasurable.Un) then show ?thesis using U by auto qed lemma measure_negligible_symdiff: assumes S: "S \ lmeasurable" and neg: "negligible (S - T \ (T - S))" shows "measure lebesgue T = measure lebesgue S" proof - have "measure lebesgue (S - T) = 0" using neg negligible_Un_eq negligible_imp_measure0 by blast then show ?thesis by (metis S Un_commute add.right_neutral lmeasurable_negligible_symdiff measure_Un2 neg negligible_Un_eq negligible_imp_measure0) qed lemma measure_closure: assumes "bounded S" and neg: "negligible (frontier S)" shows "measure lebesgue (closure S) = measure lebesgue S" proof - have "measure lebesgue (frontier S) = 0" by (metis neg negligible_imp_measure0) then show ?thesis by (metis assms lmeasurable_iff_integrable_on eq_iff_diff_eq_0 has_integral_interior integrable_on_def integral_unique lmeasurable_interior lmeasure_integral measure_frontier) qed lemma measure_interior: "\bounded S; negligible(frontier S)\ \ measure lebesgue (interior S) = measure lebesgue S" using measure_closure measure_frontier negligible_imp_measure0 by fastforce lemma measurable_Jordan: assumes "bounded S" and neg: "negligible (frontier S)" shows "S \ lmeasurable" proof - have "closure S \ lmeasurable" by (metis lmeasurable_closure \bounded S\) moreover have "interior S \ lmeasurable" by (simp add: lmeasurable_interior \bounded S\) moreover have "interior S \ S" by (simp add: interior_subset) ultimately show ?thesis using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) qed lemma measurable_convex: "\convex S; bounded S\ \ S \ lmeasurable" by (simp add: measurable_Jordan negligible_convex_frontier) lemma content_cball_conv_ball: "content (cball c r) = content (ball c r)" proof - have "ball c r - cball c r \ (cball c r - ball c r) = sphere c r" by auto hence "measure lebesgue (cball c r) = measure lebesgue (ball c r)" using negligible_sphere[of c r] by (intro measure_negligible_symdiff) simp_all thus ?thesis by simp qed subsection\Negligibility of image under non-injective linear map\ lemma negligible_Union_nat: assumes "\n::nat. negligible(S n)" shows "negligible(\n. S n)" proof - have "negligible (\m\k. S m)" for k using assms by blast then have 0: "integral UNIV (indicat_real (\m\k. S m)) = 0" and 1: "(indicat_real (\m\k. S m)) integrable_on UNIV" for k by (auto simp: negligible has_integral_iff) have 2: "\k x. indicat_real (\m\k. S m) x \ (indicat_real (\m\Suc k. S m) x)" by (simp add: indicator_def) have 3: "\x. (\k. indicat_real (\m\k. S m) x) \ (indicat_real (\n. S n) x)" by (force simp: indicator_def eventually_sequentially intro: tendsto_eventually) have 4: "bounded (range (\k. integral UNIV (indicat_real (\m\k. S m))))" by (simp add: 0) have *: "indicat_real (\n. S n) integrable_on UNIV \ (\k. integral UNIV (indicat_real (\m\k. S m))) \ (integral UNIV (indicat_real (\n. S n)))" by (intro monotone_convergence_increasing 1 2 3 4) then have "integral UNIV (indicat_real (\n. S n)) = (0::real)" using LIMSEQ_unique by (auto simp: 0) then show ?thesis using * by (simp add: negligible_UNIV has_integral_iff) qed lemma negligible_linear_singular_image: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\ inj f" shows "negligible (f ` S)" proof - obtain a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms linear_singular_image_hyperplane by blast then show "negligible (f ` S)" by (metis negligible_hyperplane negligible_subset) qed lemma measure_negligible_finite_Union: assumes "finite \" and meas: "\S. S \ \ \ S \ lmeasurable" and djointish: "pairwise (\S T. negligible (S \ T)) \" shows "measure lebesgue (\\) = (\S\\. measure lebesgue S)" using assms proof (induction) case empty then show ?case by auto next case (insert S \) then have "S \ lmeasurable" "\\ \ lmeasurable" "pairwise (\S T. negligible (S \ T)) \" by (simp_all add: fmeasurable.finite_Union insert.hyps(1) insert.prems(1) pairwise_insert subsetI) then show ?case proof (simp add: measure_Un3 insert) have *: "\T. T \ (\) S ` \ \ negligible T" using insert by (force simp: pairwise_def) have "negligible(S \ \\)" unfolding Int_Union by (rule negligible_Union) (simp_all add: * insert.hyps(1)) then show "measure lebesgue (S \ \\) = 0" using negligible_imp_measure0 by blast qed qed lemma measure_negligible_finite_Union_image: assumes "finite S" and meas: "\x. x \ S \ f x \ lmeasurable" and djointish: "pairwise (\x y. negligible (f x \ f y)) S" shows "measure lebesgue (\(f ` S)) = (\x\S. measure lebesgue (f x))" proof - have "measure lebesgue (\(f ` S)) = sum (measure lebesgue) (f ` S)" using assms by (auto simp: pairwise_mono pairwise_image intro: measure_negligible_finite_Union) also have "\ = sum (measure lebesgue \ f) S" using djointish [unfolded pairwise_def] by (metis inf.idem negligible_imp_measure0 sum.reindex_nontrivial [OF \finite S\]) also have "\ = (\x\S. measure lebesgue (f x))" by simp finally show ?thesis . qed subsection \Negligibility of a Lipschitz image of a negligible set\ text\The bound will be eliminated by a sort of onion argument\ lemma locally_Lipschitz_negl_bounded: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "0 < B" "bounded S" "negligible S" and lips: "\x. x \ S \ \T. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" unfolding negligible_iff_null_sets proof (clarsimp simp: completion.null_sets_outer) fix e::real assume "0 < e" have "S \ lmeasurable" using \negligible S\ by (simp add: negligible_iff_null_sets fmeasurableI_null_sets) then have "S \ sets lebesgue" by blast have e22: "0 < e/2 / (2 * B * real DIM('M)) ^ DIM('N)" using \0 < e\ \0 < B\ by (simp add: field_split_simps) obtain T where "open T" "S \ T" "(T - S) \ lmeasurable" "measure lebesgue (T - S) < e/2 / (2 * B * DIM('M)) ^ DIM('N)" using sets_lebesgue_outer_open [OF \S \ sets lebesgue\ e22] by (metis emeasure_eq_measure2 ennreal_leI linorder_not_le) then have T: "measure lebesgue T \ e/2 / (2 * B * DIM('M)) ^ DIM('N)" using \negligible S\ by (simp add: measure_Diff_null_set negligible_iff_null_sets) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. norm(y - x) < r \ y \ T \ (y \ S \ norm(f y - f x) \ B * norm(y - x))))" for x proof (cases "x \ S") case True obtain U where "open U" "x \ U" and U: "\y. y \ S \ U \ norm(f y - f x) \ B * norm(y - x)" using lips [OF \x \ S\] by auto have "x \ T \ U" using \S \ T\ \x \ U\ \x \ S\ by auto then obtain \ where "0 < \" "ball x \ \ T \ U" by (metis \open T\ \open U\ openE open_Int) then show ?thesis apply (rule_tac x="min (1/2) \" in exI) apply (simp del: divide_const_simps) apply (intro allI impI conjI) apply (metis dist_commute dist_norm mem_ball subsetCE) by (metis Int_iff subsetCE U dist_norm mem_ball norm_minus_commute) next case False then show ?thesis by (rule_tac x="1/4" in exI) auto qed then obtain R where R12: "\x. 0 < R x \ R x \ 1/2" and RT: "\x y. \x \ S; norm(y - x) < R x\ \ y \ T" and RB: "\x y. \x \ S; y \ S; norm(y - x) < R x\ \ norm(f y - f x) \ B * norm(y - x)" by metis+ then have gaugeR: "gauge (\x. ball x (R x))" by (simp add: gauge_def) obtain c where c: "S \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" "box (-c *\<^sub>R One:: 'M) (c *\<^sub>R One) \ {}" proof - obtain B where B: "\x. x \ S \ norm x \ B" using \bounded S\ bounded_iff by blast show ?thesis apply (rule_tac c = "abs B + 1" in that) using norm_bound_Basis_le Basis_le_norm apply (fastforce simp: box_eq_empty mem_box dest!: B intro: order_trans)+ done qed obtain \ where "countable \" and Dsub: "\\ \ cbox (-c *\<^sub>R One) (c *\<^sub>R One)" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and pw: "pairwise (\A B. interior A \ interior B = {}) \" and Ksub: "\K. K \ \ \ \x \ S \ K. K \ (\x. ball x (R x)) x" and exN: "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (2*c) / 2^n" and "S \ \\" using covering_lemma [OF c gaugeR] by force have "\u v z. K = cbox u v \ box u v \ {} \ z \ S \ z \ cbox u v \ cbox u v \ ball z (R z)" if "K \ \" for K proof - obtain u v where "K = cbox u v" using \K \ \\ cbox by blast with that show ?thesis apply (rule_tac x=u in exI) apply (rule_tac x=v in exI) apply (metis Int_iff interior_cbox cbox Ksub) done qed then obtain uf vf zf where uvz: "\K. K \ \ \ K = cbox (uf K) (vf K) \ box (uf K) (vf K) \ {} \ zf K \ S \ zf K \ cbox (uf K) (vf K) \ cbox (uf K) (vf K) \ ball (zf K) (R (zf K))" by metis define prj1 where "prj1 \ \x::'M. x \ (SOME i. i \ Basis)" define fbx where "fbx \ \D. cbox (f(zf D) - (B * DIM('M) * (prj1(vf D - uf D))) *\<^sub>R One::'N) (f(zf D) + (B * DIM('M) * prj1(vf D - uf D)) *\<^sub>R One)" have vu_pos: "0 < prj1 (vf X - uf X)" if "X \ \" for X using uvz [OF that] by (simp add: prj1_def box_ne_empty SOME_Basis inner_diff_left) have prj1_idem: "prj1 (vf X - uf X) = (vf X - uf X) \ i" if "X \ \" "i \ Basis" for X i proof - have "cbox (uf X) (vf X) \ \" using uvz \X \ \\ by auto with exN obtain n where "\i. i \ Basis \ vf X \ i - uf X \ i = (2*c) / 2^n" by blast then show ?thesis by (simp add: \i \ Basis\ SOME_Basis inner_diff prj1_def) qed have countbl: "countable (fbx ` \)" using \countable \\ by blast have "(\k\fbx`\'. measure lebesgue k) \ e/2" if "\' \ \" "finite \'" for \' proof - have BM_ge0: "0 \ B * (DIM('M) * prj1 (vf X - uf X))" if "X \ \'" for X using \0 < B\ \\' \ \\ that vu_pos by fastforce have "{} \ \'" using cbox \\' \ \\ interior_empty by blast have "(\k\fbx`\'. measure lebesgue k) \ sum (measure lebesgue o fbx) \'" by (rule sum_image_le [OF \finite \'\]) (force simp: fbx_def) also have "\ \ (\X\\'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)" proof (rule sum_mono) fix X assume "X \ \'" then have "X \ \" using \\' \ \\ by blast then have ufvf: "cbox (uf X) (vf X) = X" using uvz by blast have "prj1 (vf X - uf X) ^ DIM('M) = (\i::'M \ Basis. prj1 (vf X - uf X))" by (rule prod_constant [symmetric]) also have "\ = (\i\Basis. vf X \ i - uf X \ i)" apply (rule prod.cong [OF refl]) by (simp add: \X \ \\ inner_diff_left prj1_idem) finally have prj1_eq: "prj1 (vf X - uf X) ^ DIM('M) = (\i\Basis. vf X \ i - uf X \ i)" . have "uf X \ cbox (uf X) (vf X)" "vf X \ cbox (uf X) (vf X)" using uvz [OF \X \ \\] by (force simp: mem_box)+ moreover have "cbox (uf X) (vf X) \ ball (zf X) (1/2)" by (meson R12 order_trans subset_ball uvz [OF \X \ \\]) ultimately have "uf X \ ball (zf X) (1/2)" "vf X \ ball (zf X) (1/2)" by auto then have "dist (vf X) (uf X) \ 1" unfolding mem_ball by (metis dist_commute dist_triangle_half_l dual_order.order_iff_strict) then have 1: "prj1 (vf X - uf X) \ 1" unfolding prj1_def dist_norm using Basis_le_norm SOME_Basis order_trans by fastforce have 0: "0 \ prj1 (vf X - uf X)" using \X \ \\ prj1_def vu_pos by fastforce have "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * content (cbox (uf X) (vf X))" apply (simp add: fbx_def content_cbox_cases algebra_simps BM_ge0 \X \ \'\) apply (simp add: power_mult_distrib \0 < B\ prj1_eq [symmetric]) using MleN 0 1 uvz \X \ \\ apply (fastforce simp add: box_ne_empty power_decreasing) done also have "\ = (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" by (subst (3) ufvf[symmetric]) simp finally show "(measure lebesgue \ fbx) X \ (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" . qed also have "\ = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \'" by (simp add: sum_distrib_left) also have "\ \ e/2" proof - have div: "\' division_of \\'" apply (auto simp: \finite \'\ \{} \ \'\ division_of_def) using cbox that apply blast using pairwise_subset [OF pw \\' \ \\] unfolding pairwise_def apply force+ done have le_meaT: "measure lebesgue (\\') \ measure lebesgue T" proof (rule measure_mono_fmeasurable) show "(\\') \ sets lebesgue" using div lmeasurable_division by auto have "\\' \ \\" using \\' \ \\ by blast also have "... \ T" proof (clarify) fix x D assume "x \ D" "D \ \" show "x \ T" using Ksub [OF \D \ \\] by (metis \x \ D\ Int_iff dist_norm mem_ball norm_minus_commute subsetD RT) qed finally show "\\' \ T" . show "T \ lmeasurable" using \S \ lmeasurable\ \S \ T\ \T - S \ lmeasurable\ fmeasurable_Diff_D by blast qed have "sum (measure lebesgue) \' = sum content \'" using \\' \ \\ cbox by (force intro: sum.cong) then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \' = (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\\')" using content_division [OF div] by auto also have "\ \ (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T" apply (rule mult_left_mono [OF le_meaT]) using \0 < B\ apply (simp add: algebra_simps) done also have "\ \ e/2" using T \0 < B\ by (simp add: field_simps) finally show ?thesis . qed finally show ?thesis . qed then have e2: "sum (measure lebesgue) \ \ e/2" if "\ \ fbx ` \" "finite \" for \ by (metis finite_subset_image that) show "\W\lmeasurable. f ` S \ W \ measure lebesgue W < e" proof (intro bexI conjI) have "\X\\. f y \ fbx X" if "y \ S" for y proof - obtain X where "y \ X" "X \ \" using \S \ \\\ \y \ S\ by auto then have y: "y \ ball(zf X) (R(zf X))" using uvz by fastforce have conj_le_eq: "z - b \ y \ y \ z + b \ abs(y - z) \ b" for z y b::real by auto have yin: "y \ cbox (uf X) (vf X)" and zin: "(zf X) \ cbox (uf X) (vf X)" using uvz \X \ \\ \y \ X\ by auto have "norm (y - zf X) \ (\i\Basis. \(y - zf X) \ i\)" by (rule norm_le_l1) also have "\ \ real DIM('M) * prj1 (vf X - uf X)" proof (rule sum_bounded_above) fix j::'M assume j: "j \ Basis" show "\(y - zf X) \ j\ \ prj1 (vf X - uf X)" using yin zin j by (fastforce simp add: mem_box prj1_idem [OF \X \ \\ j] inner_diff_left) qed finally have nole: "norm (y - zf X) \ DIM('M) * prj1 (vf X - uf X)" by simp have fle: "\f y \ i - f(zf X) \ i\ \ B * DIM('M) * prj1 (vf X - uf X)" if "i \ Basis" for i proof - have "\f y \ i - f (zf X) \ i\ = \(f y - f (zf X)) \ i\" by (simp add: algebra_simps) also have "\ \ norm (f y - f (zf X))" by (simp add: Basis_le_norm that) also have "\ \ B * norm(y - zf X)" by (metis uvz RB \X \ \\ dist_commute dist_norm mem_ball \y \ S\ y) also have "\ \ B * real DIM('M) * prj1 (vf X - uf X)" using \0 < B\ by (simp add: nole) finally show ?thesis . qed show ?thesis by (rule_tac x=X in bexI) (auto simp: fbx_def prj1_idem mem_box conj_le_eq inner_add inner_diff fle \X \ \\) qed then show "f ` S \ (\D\\. fbx D)" by auto next have 1: "\D. D \ \ \ fbx D \ lmeasurable" by (auto simp: fbx_def) have 2: "I' \ \ \ finite I' \ measure lebesgue (\D\I'. fbx D) \ e/2" for I' by (rule order_trans[OF measure_Union_le e2]) (auto simp: fbx_def) show "(\D\\. fbx D) \ lmeasurable" by (intro fmeasurable_UN_bound[OF \countable \\ 1 2]) have "measure lebesgue (\D\\. fbx D) \ e/2" by (intro measure_UN_bound[OF \countable \\ 1 2]) then show "measure lebesgue (\D\\. fbx D) < e" using \0 < e\ by linarith qed qed proposition negligible_locally_Lipschitz_image: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and lips: "\x. x \ S \ \T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" shows "negligible (f ` S)" proof - let ?S = "\n. ({x \ S. norm x \ n \ (\T. open T \ x \ T \ (\y\S \ T. norm (f y - f x) \ (real n + 1) * norm (y - x)))})" have negfn: "f ` ?S n \ null_sets lebesgue" for n::nat unfolding negligible_iff_null_sets[symmetric] apply (rule_tac B = "real n + 1" in locally_Lipschitz_negl_bounded) by (auto simp: MleN bounded_iff intro: negligible_subset [OF \negligible S\]) have "S = (\n. ?S n)" proof (intro set_eqI iffI) fix x assume "x \ S" with lips obtain T B where T: "open T" "x \ T" and B: "\y. y \ S \ T \ norm(f y - f x) \ B * norm(y - x)" by metis+ have no: "norm (f y - f x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" if "y \ S \ T" for y proof - have "B * norm(y - x) \ (nat \max B (norm x)\ + 1) * norm (y - x)" by (meson max.cobounded1 mult_right_mono nat_ceiling_le_eq nat_le_iff_add norm_ge_zero order_trans) then show ?thesis using B order_trans that by blast qed have "x \ ?S (nat (ceiling (max B (norm x))))" apply (simp add: \x \ S \, rule) using real_nat_ceiling_ge max.bounded_iff apply blast using T no apply (force simp: algebra_simps) done then show "x \ (\n. ?S n)" by force qed auto then show ?thesis by (rule ssubst) (auto simp: image_Union negligible_iff_null_sets intro: negfn) qed corollary negligible_differentiable_image_negligible: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MleN: "DIM('M) \ DIM('N)" "negligible S" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "\T B. open T \ x \ T \ (\y \ S \ T. norm(f y - f x) \ B * norm(y - x))" if "x \ S" for x proof - obtain f' where "linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x)" using diff_f \x \ S\ by (auto simp: linear_linear differentiable_on_def differentiable_def has_derivative_within_alt) obtain B where "B > 0" and B: "\x. norm (f' x) \ B * norm x" using linear_bounded_pos \linear f'\ by blast obtain d where "d>0" and d: "\y. \y \ S; norm (y - x) < d\ \ norm (f y - f x - f' (y - x)) \ norm (y - x)" using f' [of 1] by (force simp:) have *: "norm (f y - f x) \ (B + 1) * norm (y - x)" if "y \ S" "norm (y - x) < d" for y proof - have "norm (f y - f x) -B * norm (y - x) \ norm (f y - f x) - norm (f' (y - x))" by (simp add: B) also have "\ \ norm (f y - f x - f' (y - x))" by (rule norm_triangle_ineq2) also have "... \ norm (y - x)" by (rule d [OF that]) finally show ?thesis by (simp add: algebra_simps) qed show ?thesis apply (rule_tac x="ball x d" in exI) apply (rule_tac x="B+1" in exI) using \d>0\ apply (auto simp: dist_norm norm_minus_commute intro!: *) done qed with negligible_locally_Lipschitz_image assms show ?thesis by metis qed corollary negligible_differentiable_image_lowdim: fixes f :: "'M::euclidean_space \ 'N::euclidean_space" assumes MlessN: "DIM('M) < DIM('N)" and diff_f: "f differentiable_on S" shows "negligible (f ` S)" proof - have "x \ DIM('M) \ x \ DIM('N)" for x using MlessN by linarith obtain lift :: "'M * real \ 'N" and drop :: "'N \ 'M * real" and j :: 'N where "linear lift" "linear drop" and dropl [simp]: "\z. drop (lift z) = z" and "j \ Basis" and j: "\x. lift(x,0) \ j = 0" using lowerdim_embeddings [OF MlessN] by metis have "negligible {x. x\j = 0}" by (metis \j \ Basis\ negligible_standard_hyperplane) then have neg0S: "negligible ((\x. lift (x, 0)) ` S)" apply (rule negligible_subset) by (simp add: image_subsetI j) have diff_f': "f \ fst \ drop differentiable_on (\x. lift (x, 0)) ` S" using diff_f apply (clarsimp simp add: differentiable_on_def) apply (intro differentiable_chain_within linear_imp_differentiable [OF \linear drop\] - linear_imp_differentiable [OF fst_linear]) + linear_imp_differentiable [OF linear_fst]) apply (force simp: image_comp o_def) done have "f = (f o fst o drop o (\x. lift (x, 0)))" by (simp add: o_def) then show ?thesis apply (rule ssubst) apply (subst image_comp [symmetric]) apply (metis negligible_differentiable_image_negligible order_refl diff_f' neg0S) done qed subsection\Measurability of countable unions and intersections of various kinds.\ lemma assumes S: "\n. S n \ lmeasurable" and leB: "\n. measure lebesgue (S n) \ B" and nest: "\n. S n \ S(Suc n)" shows measurable_nested_Union: "(\n. S n) \ lmeasurable" and measure_nested_Union: "(\n. measure lebesgue (S n)) \ measure lebesgue (\n. S n)" (is ?Lim) proof - have 1: "\n. (indicat_real (S n)) integrable_on UNIV" using S measurable_integrable by blast have 2: "\n x::'a. indicat_real (S n) x \ (indicat_real (S (Suc n)) x)" by (simp add: indicator_leI nest rev_subsetD) have 3: "\x. (\n. indicat_real (S n) x) \ (indicat_real (\(S ` UNIV)) x)" apply (rule tendsto_eventually) apply (simp add: indicator_def) by (metis eventually_sequentiallyI lift_Suc_mono_le nest subsetCE) have 4: "bounded (range (\n. integral UNIV (indicat_real (S n))))" using leB by (auto simp: lmeasure_integral_UNIV [symmetric] S bounded_iff) have "(\n. S n) \ lmeasurable \ ?Lim" apply (simp add: lmeasure_integral_UNIV S cong: conj_cong) apply (simp add: measurable_integrable) apply (rule monotone_convergence_increasing [OF 1 2 3 4]) done then show "(\n. S n) \ lmeasurable" "?Lim" by auto qed lemma assumes S: "\n. S n \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and leB: "\n. (\k\n. measure lebesgue (S k)) \ B" shows measurable_countable_negligible_Union: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - have 1: "\ (S ` {..n}) \ lmeasurable" for n using S by blast have 2: "measure lebesgue (\ (S ` {..n})) \ B" for n proof - have "measure lebesgue (\ (S ` {..n})) \ (\k\n. measure lebesgue (S k))" by (simp add: S fmeasurableD measure_UNION_le) with leB show ?thesis using order_trans by blast qed have 3: "\n. \ (S ` {..n}) \ \ (S ` {..Suc n})" by (simp add: SUP_subset_mono) have eqS: "(\n. S n) = (\n. \ (S ` {..n}))" using atLeastAtMost_iff by blast also have "(\n. \ (S ` {..n})) \ lmeasurable" by (intro measurable_nested_Union [OF 1 2] 3) finally show "(\n. S n) \ lmeasurable" . have eqm: "(\i\n. measure lebesgue (S i)) = measure lebesgue (\ (S ` {..n}))" for n using assms by (simp add: measure_negligible_finite_Union_image pairwise_mono) have "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. \ (S ` {..n}))" by (simp add: sums_def' eqm atLeast0AtMost) (intro measure_nested_Union [OF 1 2] 3) then show ?Sums by (simp add: eqS) qed lemma negligible_countable_Union [intro]: assumes "countable \" and meas: "\S. S \ \ \ negligible S" shows "negligible (\\)" proof (cases "\ = {}") case False then show ?thesis by (metis from_nat_into range_from_nat_into assms negligible_Union_nat) qed simp lemma assumes S: "\n. (S n) \ lmeasurable" and djointish: "pairwise (\m n. negligible (S m \ S n)) UNIV" and bo: "bounded (\n. S n)" shows measurable_countable_negligible_Union_bounded: "(\n. S n) \ lmeasurable" and measure_countable_negligible_Union_bounded: "(\n. (measure lebesgue (S n))) sums measure lebesgue (\n. S n)" (is ?Sums) proof - obtain a b where ab: "(\n. S n) \ cbox a b" using bo bounded_subset_cbox_symmetric by metis then have B: "(\k\n. measure lebesgue (S k)) \ measure lebesgue (cbox a b)" for n proof - have "(\k\n. measure lebesgue (S k)) = measure lebesgue (\ (S ` {..n}))" using measure_negligible_finite_Union_image [OF _ _ pairwise_subset] djointish by (metis S finite_atMost subset_UNIV) also have "\ \ measure lebesgue (cbox a b)" apply (rule measure_mono_fmeasurable) using ab S by force+ finally show ?thesis . qed show "(\n. S n) \ lmeasurable" by (rule measurable_countable_negligible_Union [OF S djointish B]) show ?Sums by (rule measure_countable_negligible_Union [OF S djointish B]) qed lemma measure_countable_Union_approachable: assumes "countable \" "e > 0" and measD: "\d. d \ \ \ d \ lmeasurable" and B: "\D'. \D' \ \; finite D'\ \ measure lebesgue (\D') \ B" obtains D' where "D' \ \" "finite D'" "measure lebesgue (\\) - e < measure lebesgue (\D')" proof (cases "\ = {}") case True then show ?thesis by (simp add: \e > 0\ that) next case False let ?S = "\n. \k \ n. from_nat_into \ k" have "(\n. measure lebesgue (?S n)) \ measure lebesgue (\n. ?S n)" proof (rule measure_nested_Union) show "?S n \ lmeasurable" for n by (simp add: False fmeasurable.finite_UN from_nat_into measD) show "measure lebesgue (?S n) \ B" for n by (metis (mono_tags, lifting) B False finite_atMost finite_imageI from_nat_into image_iff subsetI) show "?S n \ ?S (Suc n)" for n by force qed then obtain N where N: "\n. n \ N \ dist (measure lebesgue (?S n)) (measure lebesgue (\n. ?S n)) < e" using metric_LIMSEQ_D \e > 0\ by blast show ?thesis proof show "from_nat_into \ ` {..N} \ \" by (auto simp: False from_nat_into) have eq: "(\n. \k\n. from_nat_into \ k) = (\\)" using \countable \\ False by (auto intro: from_nat_into dest: from_nat_into_surj [OF \countable \\]) show "measure lebesgue (\\) - e < measure lebesgue (\ (from_nat_into \ ` {..N}))" using N [OF order_refl] by (auto simp: eq algebra_simps dist_norm) qed auto qed subsection\Negligibility is a local property\ lemma locally_negligible_alt: "negligible S \ (\x \ S. \U. openin (top_of_set S) U \ x \ U \ negligible U)" (is "_ = ?rhs") proof assume "negligible S" then show ?rhs using openin_subtopology_self by blast next assume ?rhs then obtain U where ope: "\x. x \ S \ openin (top_of_set S) (U x)" and cov: "\x. x \ S \ x \ U x" and neg: "\x. x \ S \ negligible (U x)" by metis obtain \ where "\ \ U ` S" "countable \" and eq: "\\ = \(U ` S)" using ope by (force intro: Lindelof_openin [of "U ` S" S]) then have "negligible (\\)" by (metis imageE neg negligible_countable_Union subset_eq) with eq have "negligible (\(U ` S))" by metis moreover have "S \ \(U ` S)" using cov by blast ultimately show "negligible S" using negligible_subset by blast qed lemma locally_negligible: "locally negligible S \ negligible S" unfolding locally_def apply safe apply (meson negligible_subset openin_subtopology_self locally_negligible_alt) by (meson negligible_subset openin_imp_subset order_refl) subsection\Integral bounds\ lemma set_integral_norm_bound: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "set_integrable M k f \ norm (LINT x:k|M. f x) \ LINT x:k|M. norm (f x)" using integral_norm_bound[of M "\x. indicator k x *\<^sub>R f x"] by (simp add: set_lebesgue_integral_def) lemma set_integral_finite_UN_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" and ae: "\i j. i \ I \ j \ I \ AE x in M. (x \ A i \ x \ A j) \ i = j" and [measurable]: "\i. i \ I \ A i \ sets M" and f: "\i. i \ I \ set_integrable M (A i) f" shows "LINT x:(\i\I. A i)|M. f x = (\i\I. LINT x:A i|M. f x)" using \finite I\ order_refl[of I] proof (induction I rule: finite_subset_induct') case (insert i I') have "AE x in M. (\j\I'. x \ A i \ x \ A j)" proof (intro AE_ball_countable[THEN iffD2] ballI) fix j assume "j \ I'" with \I' \ I\ \i \ I'\ have "i \ j" "j \ I" by auto then show "AE x in M. x \ A i \ x \ A j" using ae[of i j] \i \ I\ by auto qed (use \finite I'\ in \rule countable_finite\) then have "AE x\A i in M. \xa\I'. x \ A xa " by auto with insert.hyps insert.IH[symmetric] show ?case by (auto intro!: set_integral_Un_AE sets.finite_UN f set_integrable_UN) qed (simp add: set_lebesgue_integral_def) lemma set_integrable_norm: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M k f" shows "set_integrable M k (\x. norm (f x))" using integrable_norm f by (force simp add: set_integrable_def) lemma absolutely_integrable_bounded_variation: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f absolutely_integrable_on UNIV" obtains B where "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" proof (rule that[of "integral UNIV (\x. norm (f x))"]; safe) fix d :: "'a set set" assume d: "d division_of \d" have *: "k \ d \ f absolutely_integrable_on k" for k using f[THEN set_integrable_subset, of k] division_ofD(2,4)[OF d, of k] by auto note d' = division_ofD[OF d] have "(\k\d. norm (integral k f)) = (\k\d. norm (LINT x:k|lebesgue. f x))" by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *) also have "\ \ (\k\d. LINT x:k|lebesgue. norm (f x))" by (intro sum_mono set_integral_norm_bound *) also have "\ = (\k\d. integral k (\x. norm (f x)))" by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *) also have "\ \ integral (\d) (\x. norm (f x))" using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def by (subst integral_combine_division_topdown[OF _ d]) auto also have "\ \ integral UNIV (\x. norm (f x))" using integrable_on_subdivision[OF d] assms unfolding absolutely_integrable_on_def by (intro integral_subset_le) auto finally show "(\k\d. norm (integral k f)) \ integral UNIV (\x. norm (f x))" . qed lemma absdiff_norm_less: assumes "sum (\x. norm (f x - g x)) s < e" and "finite s" shows "\sum (\x. norm(f x)) s - sum (\x. norm(g x)) s\ < e" unfolding sum_subtractf[symmetric] apply (rule le_less_trans[OF sum_abs]) apply (rule le_less_trans[OF _ assms(1)]) apply (rule sum_mono) apply (rule norm_triangle_ineq3) done proposition bounded_variation_absolutely_integrable_interval: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes f: "f integrable_on cbox a b" and *: "\d. d division_of (cbox a b) \ sum (\K. norm(integral K f)) d \ B" shows "f absolutely_integrable_on cbox a b" proof - let ?f = "\d. \K\d. norm (integral K f)" and ?D = "{d. d division_of (cbox a b)}" have D_1: "?D \ {}" by (rule elementary_interval[of a b]) auto have D_2: "bdd_above (?f`?D)" by (metis * mem_Collect_eq bdd_aboveI2) note D = D_1 D_2 let ?S = "SUP x\?D. ?f x" have *: "\\. gauge \ \ (\p. p tagged_division_of cbox a b \ \ fine p \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e)" if e: "e > 0" for e proof - have "?S - e/2 < ?S" using \e > 0\ by simp then obtain d where d: "d division_of (cbox a b)" "?S - e/2 < (\k\d. norm (integral k f))" unfolding less_cSUP_iff[OF D] by auto note d' = division_ofD[OF this(1)] have "\e>0. \i\d. x \ i \ ball x e \ i = {}" for x proof - have "\d'>0. \x'\\{i \ d. x \ i}. d' \ dist x x'" proof (rule separate_point_closed) show "closed (\{i \ d. x \ i})" using d' by force show "x \ \{i \ d. x \ i}" by auto qed then show ?thesis by force qed then obtain k where k: "\x. 0 < k x" "\i x. \i \ d; x \ i\ \ ball x (k x) \ i = {}" by metis have "e/2 > 0" using e by auto with Henstock_lemma[OF f] obtain \ where g: "gauge \" "\p. \p tagged_partial_division_of cbox a b; \ fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (metis (no_types, lifting)) let ?g = "\x. \ x \ ball x (k x)" show ?thesis proof (intro exI conjI allI impI) show "gauge ?g" using g(1) k(1) by (auto simp: gauge_def) next fix p assume "p tagged_division_of (cbox a b) \ ?g fine p" then have p: "p tagged_division_of cbox a b" "\ fine p" "(\x. ball x (k x)) fine p" by (auto simp: fine_Int) note p' = tagged_division_ofD[OF p(1)] define p' where "p' = {(x,k) | x k. \i l. x \ i \ i \ d \ (x,l) \ p \ k = i \ l}" have gp': "\ fine p'" using p(2) by (auto simp: p'_def fine_def) have p'': "p' tagged_division_of (cbox a b)" proof (rule tagged_division_ofI) show "finite p'" proof (rule finite_subset) show "p' \ (\(k, x, l). (x, k \ l)) ` (d \ p)" by (force simp: p'_def image_iff) show "finite ((\(k, x, l). (x, k \ l)) ` (d \ p))" by (simp add: d'(1) p'(1)) qed next fix x K assume "(x, K) \ p'" then have "\i l. x \ i \ i \ d \ (x, l) \ p \ K = i \ l" unfolding p'_def by auto then obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" by blast show "x \ K" and "K \ cbox a b" using p'(2-3)[OF il(3)] il by auto show "\a b. K = cbox a b" unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)] by (meson Int_interval) next fix x1 K1 assume "(x1, K1) \ p'" then have "\i l. x1 \ i \ i \ d \ (x1, l) \ p \ K1 = i \ l" unfolding p'_def by auto then obtain i1 l1 where il1: "x1 \ i1" "i1 \ d" "(x1, l1) \ p" "K1 = i1 \ l1" by blast fix x2 K2 assume "(x2,K2) \ p'" then have "\i l. x2 \ i \ i \ d \ (x2, l) \ p \ K2 = i \ l" unfolding p'_def by auto then obtain i2 l2 where il2: "x2 \ i2" "i2 \ d" "(x2, l2) \ p" "K2 = i2 \ l2" by blast assume "(x1, K1) \ (x2, K2)" then have "interior i1 \ interior i2 = {} \ interior l1 \ interior l2 = {}" using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)] by (auto simp: il1 il2) then show "interior K1 \ interior K2 = {}" unfolding il1 il2 by auto next have *: "\(x, X) \ p'. X \ cbox a b" unfolding p'_def using d' by blast have "y \ \{K. \x. (x, K) \ p'}" if y: "y \ cbox a b" for y proof - obtain x l where xl: "(x, l) \ p" "y \ l" using y unfolding p'(6)[symmetric] by auto obtain i where i: "i \ d" "y \ i" using y unfolding d'(6)[symmetric] by auto have "x \ i" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show ?thesis unfolding p'_def by (rule_tac X="i \ l" in UnionI) (use i xl in auto) qed show "\{K. \x. (x, K) \ p'} = cbox a b" proof show "\{k. \x. (x, k) \ p'} \ cbox a b" using * by auto next show "cbox a b \ \{k. \x. (x, k) \ p'}" proof fix y assume y: "y \ cbox a b" obtain x L where xl: "(x, L) \ p" "y \ L" using y unfolding p'(6)[symmetric] by auto obtain I where i: "I \ d" "y \ I" using y unfolding d'(6)[symmetric] by auto have "x \ I" using fineD[OF p(3) xl(1)] using k(2) i xl by auto then show "y \ \{k. \x. (x, k) \ p'}" apply (rule_tac X="I \ L" in UnionI) using i xl by (auto simp: p'_def) qed qed qed then have sum_less_e2: "(\(x,K) \ p'. norm (content K *\<^sub>R f x - integral K f)) < e/2" using g(2) gp' tagged_division_of_def by blast have "(x, I \ L) \ p'" if x: "(x, L) \ p" "I \ d" and y: "y \ I" "y \ L" for x I L y proof - have "x \ I" using fineD[OF p(3) that(1)] k(2)[OF \I \ d\] y by auto with x have "(\i l. x \ i \ i \ d \ (x, l) \ p \ I \ L = i \ l)" by blast then have "(x, I \ L) \ p'" by (simp add: p'_def) with y show ?thesis by auto qed moreover have "\y i l. (x, K) = (y, i \ l) \ (y, l) \ p \ i \ d \ i \ l \ {}" if xK: "(x,K) \ p'" for x K proof - obtain i l where il: "x \ i" "i \ d" "(x, l) \ p" "K = i \ l" using xK unfolding p'_def by auto then show ?thesis using p'(2) by fastforce qed ultimately have p'alt: "p' = {(x, I \ L) | x I L. (x,L) \ p \ I \ d \ I \ L \ {}}" by auto have sum_p': "(\(x,K) \ p'. norm (integral K f)) = (\k\snd ` p'. norm (integral k f))" apply (subst sum.over_tagged_division_lemma[OF p'',of "\k. norm (integral k f)"]) apply (auto intro: integral_null simp: content_eq_0_interior) done have snd_p_div: "snd ` p division_of cbox a b" by (rule division_of_tagged_division[OF p(1)]) note snd_p = division_ofD[OF snd_p_div] have fin_d_sndp: "finite (d \ snd ` p)" by (simp add: d'(1) snd_p(1)) have *: "\sni sni' sf sf'. \\sf' - sni'\ < e/2; ?S - e/2 < sni; sni' \ ?S; sni \ sni'; sf' = sf\ \ \sf - ?S\ < e" by arith show "norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - ?S) < e" unfolding real_norm_def proof (rule *) show "\(\(x,K)\p'. norm (content K *\<^sub>R f x)) - (\(x,k)\p'. norm (integral k f))\ < e/2" using p'' sum_less_e2 unfolding split_def by (force intro!: absdiff_norm_less) show "(\(x,k) \ p'. norm (integral k f)) \?S" by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper) show "(\k\d. norm (integral k f)) \ (\(x,k) \ p'. norm (integral k f))" proof - have *: "{k \ l | k l. k \ d \ l \ snd ` p} = (\(k,l). k \ l) ` (d \ snd ` p)" by auto have "(\K\d. norm (integral K f)) \ (\i\d. \l\snd ` p. norm (integral (i \ l) f))" proof (rule sum_mono) fix K assume k: "K \ d" from d'(4)[OF this] obtain u v where uv: "K = cbox u v" by metis define d' where "d' = {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" have uvab: "cbox u v \ cbox a b" using d(1) k uv by blast have "d' division_of cbox u v" unfolding d'_def by (rule division_inter_1 [OF snd_p_div uvab]) moreover then have "norm (\i\d'. integral i f) \ (\k\d'. norm (integral k f))" by (simp add: sum_norm_le) ultimately have "norm (integral K f) \ sum (\k. norm (integral k f)) d'" apply (subst integral_combine_division_topdown[of _ _ d']) apply (auto simp: uv intro: integrable_on_subcbox[OF assms(1) uvab]) done also have "\ = (\I\{K \ L |L. L \ snd ` p}. norm (integral I f))" proof - have *: "norm (integral I f) = 0" if "I \ {cbox u v \ l |l. l \ snd ` p}" "I \ {cbox u v \ l |l. l \ snd ` p \ cbox u v \ l \ {}}" for I using that by auto show ?thesis apply (rule sum.mono_neutral_left) apply (simp add: snd_p(1)) unfolding d'_def uv using * by auto qed also have "\ = (\l\snd ` p. norm (integral (K \ l) f))" proof - have *: "norm (integral (K \ l) f) = 0" if "l \ snd ` p" "y \ snd ` p" "l \ y" "K \ l = K \ y" for l y proof - have "interior (K \ l) \ interior (l \ y)" by (metis Int_lower2 interior_mono le_inf_iff that(4)) then have "interior (K \ l) = {}" by (simp add: snd_p(5) that) moreover from d'(4)[OF k] snd_p(4)[OF that(1)] obtain u1 v1 u2 v2 where uv: "K = cbox u1 u2" "l = cbox v1 v2" by metis ultimately show ?thesis using that integral_null unfolding uv Int_interval content_eq_0_interior by (metis (mono_tags, lifting) norm_eq_zero) qed show ?thesis unfolding Setcompr_eq_image apply (rule sum.reindex_nontrivial [unfolded o_def]) apply (rule finite_imageI) apply (rule p') using * by auto qed finally show "norm (integral K f) \ (\l\snd ` p. norm (integral (K \ l) f))" . qed also have "\ = (\(i,l) \ d \ snd ` p. norm (integral (i\l) f))" by (simp add: sum.cartesian_product) also have "\ = (\x \ d \ snd ` p. norm (integral (case_prod (\) x) f))" by (force simp: split_def intro!: sum.cong) also have "\ = (\k\{i \ l |i l. i \ d \ l \ snd ` p}. norm (integral k f))" proof - have eq0: " (integral (l1 \ k1) f) = 0" if "l1 \ k1 = l2 \ k2" "(l1, k1) \ (l2, k2)" "l1 \ d" "(j1,k1) \ p" "l2 \ d" "(j2,k2) \ p" for l1 l2 k1 k2 j1 j2 proof - obtain u1 v1 u2 v2 where uv: "l1 = cbox u1 u2" "k1 = cbox v1 v2" using \(j1, k1) \ p\ \l1 \ d\ d'(4) p'(4) by blast have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" by (meson d'(5) old.prod.inject p'(5) that(3) that(4) that(5) that(6)) moreover have "interior (l1 \ k1) = interior (l2 \ k2)" by (simp add: that(1)) ultimately have "interior(l1 \ k1) = {}" by auto then show ?thesis unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed show ?thesis unfolding * apply (rule sum.reindex_nontrivial [OF fin_d_sndp, symmetric, unfolded o_def]) apply clarsimp by (metis eq0 fst_conv snd_conv) qed also have "\ = (\(x,k) \ p'. norm (integral k f))" proof - have 0: "integral (ia \ snd (a, b)) f = 0" if "ia \ snd (a, b) \ snd ` p'" "ia \ d" "(a, b) \ p" for ia a b proof - have "ia \ b = {}" using that unfolding p'alt image_iff Bex_def not_ex apply (erule_tac x="(a, ia \ b)" in allE) apply auto done then show ?thesis by auto qed have 1: "\i l. snd (a, b) = i \ l \ i \ d \ l \ snd ` p" if "(a, b) \ p'" for a b using that apply (clarsimp simp: p'_def image_iff) by (metis (no_types, hide_lams) snd_conv) show ?thesis unfolding sum_p' apply (rule sum.mono_neutral_right) apply (metis * finite_imageI[OF fin_d_sndp]) using 0 1 by auto qed finally show ?thesis . qed show "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - let ?S = "{(x, i \ l) |x i l. (x, l) \ p \ i \ d}" have *: "?S = (\(xl,i). (fst xl, snd xl \ i)) ` (p \ d)" by force have fin_pd: "finite (p \ d)" using finite_cartesian_product[OF p'(1) d'(1)] by metis have "(\(x,k) \ p'. norm (content k *\<^sub>R f x)) = (\(x,k) \ ?S. \content k\ * norm (f x))" unfolding norm_scaleR apply (rule sum.mono_neutral_left) apply (subst *) apply (rule finite_imageI [OF fin_pd]) unfolding p'alt apply auto by fastforce also have "\ = (\((x,l),i)\p \ d. \content (l \ i)\ * norm (f x))" proof - have "\content (l1 \ k1)\ * norm (f x1) = 0" if "(x1, l1) \ p" "(x2, l2) \ p" "k1 \ d" "k2 \ d" "x1 = x2" "l1 \ k1 = l2 \ k2" "x1 \ x2 \ l1 \ l2 \ k1 \ k2" for x1 l1 k1 x2 l2 k2 proof - obtain u1 v1 u2 v2 where uv: "k1 = cbox u1 u2" "l1 = cbox v1 v2" by (meson \(x1, l1) \ p\ \k1 \ d\ d(1) division_ofD(4) p'(4)) have "l1 \ l2 \ k1 \ k2" using that by auto then have "interior k1 \ interior k2 = {} \ interior l1 \ interior l2 = {}" apply (rule disjE) using that p'(5) d'(5) by auto moreover have "interior (l1 \ k1) = interior (l2 \ k2)" unfolding that .. ultimately have "interior (l1 \ k1) = {}" by auto then show "\content (l1 \ k1)\ * norm (f x1) = 0" unfolding uv Int_interval content_eq_0_interior[symmetric] by auto qed then show ?thesis unfolding * apply (subst sum.reindex_nontrivial [OF fin_pd]) unfolding split_paired_all o_def split_def prod.inject apply force+ done qed also have "\ = (\(x,k) \ p. content k *\<^sub>R norm (f x))" proof - have sumeq: "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" if "(x, l) \ p" for x l proof - note xl = p'(2-4)[OF that] then obtain u v where uv: "l = cbox u v" by blast have "(\i\d. \content (l \ i)\) = (\k\d. content (k \ cbox u v))" by (simp add: Int_commute uv) also have "\ = sum content {k \ cbox u v| k. k \ d}" proof - have eq0: "content (k \ cbox u v) = 0" if "k \ d" "y \ d" "k \ y" and eq: "k \ cbox u v = y \ cbox u v" for k y proof - from d'(4)[OF that(1)] d'(4)[OF that(2)] obtain \ \ where \: "k \ cbox u v = cbox \ \" by (meson Int_interval) have "{} = interior ((k \ y) \ cbox u v)" by (simp add: d'(5) that) also have "\ = interior (y \ (k \ cbox u v))" by auto also have "\ = interior (k \ cbox u v)" unfolding eq by auto finally show ?thesis unfolding \ content_eq_0_interior .. qed then show ?thesis unfolding Setcompr_eq_image apply (rule sum.reindex_nontrivial [OF \finite d\, unfolded o_def, symmetric]) by auto qed also have "\ = sum content {cbox u v \ k |k. k \ d \ cbox u v \ k \ {}}" apply (rule sum.mono_neutral_right) unfolding Setcompr_eq_image apply (rule finite_imageI [OF \finite d\]) apply (fastforce simp: inf.commute)+ done finally show "(\i\d. content (l \ i) * norm (f x)) = content l * norm (f x)" unfolding sum_distrib_right[symmetric] real_scaleR_def apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]]) using xl(2)[unfolded uv] unfolding uv apply auto done qed show ?thesis by (subst sum_Sigma_product[symmetric]) (auto intro!: sumeq sum.cong p' d') qed finally show ?thesis . qed qed (rule d) qed qed then show ?thesis using absolutely_integrable_onI [OF f has_integral_integrable] has_integral[of _ ?S] by blast qed lemma bounded_variation_absolutely_integrable: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on UNIV" and "\d. d division_of (\d) \ sum (\k. norm (integral k f)) d \ B" shows "f absolutely_integrable_on UNIV" proof (rule absolutely_integrable_onI, fact) let ?f = "\d. \k\d. norm (integral k f)" and ?D = "{d. d division_of (\d)}" have D_1: "?D \ {}" by (rule elementary_interval) auto have D_2: "bdd_above (?f`?D)" by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp note D = D_1 D_2 let ?S = "SUP d\?D. ?f d" have "\a b. f integrable_on cbox a b" using assms(1) integrable_on_subcbox by blast then have f_int: "\a b. f absolutely_integrable_on cbox a b" apply (rule bounded_variation_absolutely_integrable_interval[where B=B]) using assms(2) apply blast done have "((\x. norm (f x)) has_integral ?S) UNIV" apply (subst has_integral_alt') apply safe proof goal_cases case (1 a b) show ?case using f_int[of a b] unfolding absolutely_integrable_on_def by auto next case prems: (2 e) have "\y\sum (\k. norm (integral k f)) ` {d. d division_of \d}. \ y \ ?S - e" proof (rule ccontr) assume "\ ?thesis" then have "?S \ ?S - e" by (intro cSUP_least[OF D(1)]) auto then show False using prems by auto qed then obtain d K where ddiv: "d division_of \d" and "K = (\k\d. norm (integral k f))" "Sup (sum (\k. norm (integral k f)) ` {d. d division_of \ d}) - e < K" by (auto simp add: image_iff not_le) then have d: "Sup (sum (\k. norm (integral k f)) ` {d. d division_of \ d}) - e < (\k\d. norm (integral k f))" by auto note d'=division_ofD[OF ddiv] have "bounded (\d)" by (rule elementary_bounded,fact) from this[unfolded bounded_pos] obtain K where K: "0 < K" "\x\\d. norm x \ K" by auto show ?case proof (intro conjI impI allI exI) fix a b :: 'n assume ab: "ball 0 (K + 1) \ cbox a b" have *: "\s s1. \?S - e < s1; s1 \ s; s < ?S + e\ \ \s - ?S\ < e" by arith show "norm (integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) - ?S) < e" unfolding real_norm_def proof (rule * [OF d]) have "(\k\d. norm (integral k f)) \ sum (\k. integral k (\x. norm (f x))) d" proof (intro sum_mono) fix k assume "k \ d" with d'(4) f_int show "norm (integral k f) \ integral k (\x. norm (f x))" by (force simp: absolutely_integrable_on_def integral_norm_bound_integral) qed also have "\ = integral (\d) (\x. norm (f x))" apply (rule integral_combine_division_bottomup[OF ddiv, symmetric]) using absolutely_integrable_on_def d'(4) f_int by blast also have "\ \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" proof - have "\d \ cbox a b" using K(2) ab by fastforce then show ?thesis using integrable_on_subdivision[OF ddiv] f_int[of a b] unfolding absolutely_integrable_on_def by (auto intro!: integral_subset_le) qed finally show "(\k\d. norm (integral k f)) \ integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0)" . next have "e/2>0" using \e > 0\ by auto moreover have f: "f integrable_on cbox a b" "(\x. norm (f x)) integrable_on cbox a b" using f_int by (auto simp: absolutely_integrable_on_def) ultimately obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of (cbox a b); d1 fine p\ \ norm ((\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm (f x))) < e/2" unfolding has_integral_integral has_integral by meson obtain d2 where "gauge d2" and d2: "\p. \p tagged_partial_division_of (cbox a b); d2 fine p\ \ (\(x,k) \ p. norm (content k *\<^sub>R f x - integral k f)) < e/2" by (blast intro: Henstock_lemma [OF f(1) \e/2>0\]) obtain p where p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p" by (rule fine_division_exists [OF gauge_Int [OF \gauge d1\ \gauge d2\], of a b]) (auto simp add: fine_Int) have *: "\sf sf' si di. \sf' = sf; si \ ?S; \sf - si\ < e/2; \sf' - di\ < e/2\ \ di < ?S + e" by arith have "integral (cbox a b) (\x. norm (f x)) < ?S + e" proof (rule *) show "\(\(x,k)\p. norm (content k *\<^sub>R f x)) - (\(x,k)\p. norm (integral k f))\ < e/2" unfolding split_def apply (rule absdiff_norm_less) using d2[of p] p(1,3) apply (auto simp: tagged_division_of_def split_def) done show "\(\(x,k) \ p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\x. norm(f x))\ < e/2" using d1[OF p(1,2)] by (simp only: real_norm_def) show "(\(x,k) \ p. content k *\<^sub>R norm (f x)) = (\(x,k) \ p. norm (content k *\<^sub>R f x))" by (auto simp: split_paired_all sum.cong [OF refl]) show "(\(x,k) \ p. norm (integral k f)) \ ?S" using partial_division_of_tagged_division[of p "cbox a b"] p(1) apply (subst sum.over_tagged_division_lemma[OF p(1)]) apply (auto simp: content_eq_0_interior tagged_partial_division_of_def intro!: cSUP_upper2 D) done qed then show "integral (cbox a b) (\x. if x \ UNIV then norm (f x) else 0) < ?S + e" by simp qed qed (insert K, auto) qed then show "(\x. norm (f x)) integrable_on UNIV" by blast qed subsection\Outer and inner approximation of measurable sets by well-behaved sets.\ proposition measurable_outer_intervals_bounded: assumes "S \ lmeasurable" "S \ cbox a b" "e > 0" obtains \ where "countable \" "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" "pairwise (\A B. interior A \ interior B = {}) \" "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i)/2^n" "\K. \K \ \; box a b \ {}\ \ interior K \ {}" "S \ \\" "\\ \ lmeasurable" "measure lebesgue (\\) \ measure lebesgue S + e" proof (cases "box a b = {}") case True show ?thesis proof (cases "cbox a b = {}") case True with assms have [simp]: "S = {}" by auto show ?thesis proof show "countable {}" by simp qed (use \e > 0\ in auto) next case False show ?thesis proof show "countable {cbox a b}" by simp show "\u v. cbox u v \ {cbox a b} \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i)/2 ^ n" using False by (force simp: eq_cbox intro: exI [where x=0]) show "measure lebesgue (\{cbox a b}) \ measure lebesgue S + e" using assms by (simp add: sum_content.box_empty_imp [OF True]) qed (use assms \cbox a b \ {}\ in auto) qed next case False let ?\ = "measure lebesgue" have "S \ cbox a b \ lmeasurable" using \S \ lmeasurable\ by blast then have indS_int: "(indicator S has_integral (?\ S)) (cbox a b)" by (metis integral_indicator \S \ cbox a b\ has_integral_integrable_integral inf.orderE integrable_on_indicator) with \e > 0\ obtain \ where "gauge \" and \: "\\. \\ tagged_division_of (cbox a b); \ fine \\ \ norm ((\(x,K)\\. content(K) *\<^sub>R indicator S x) - ?\ S) < e" by (force simp: has_integral) have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) obtain \ where \: "countable \" "\\ \ cbox a b" and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ \ x" and close: "\u v. cbox u v \ \ \ \n. \i \ Basis. v\i - u\i = (b\i - a\i)/2^n" and covers: "S \ \\" using covering_lemma [of S a b \] \gauge \\ \box a b \ {}\ assms by force show ?thesis proof show "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" by (meson Sup_le_iff \(2) cbox interior_empty) have negl_int: "negligible(K \ L)" if "K \ \" "L \ \" "K \ L" for K L proof - have "interior K \ interior L = {}" using djointish pairwiseD that by fastforce moreover obtain u v x y where "K = cbox u v" "L = cbox x y" using cbox \K \ \\ \L \ \\ by blast ultimately show ?thesis by (simp add: Int_interval box_Int_box negligible_interval(1)) qed have fincase: "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" if "finite \" "\ \ \" for \ proof - obtain t where t: "\K. K \ \ \ t K \ S \ K \ K \ \(t K)" using covered \\ \ \\ subsetD by metis have "\K \ \. \L \ \. K \ L \ interior K \ interior L = {}" using that djointish by (simp add: pairwise_def) (metis subsetD) with cbox that \ have \div: "\ division_of (\\)" by (fastforce simp: division_of_def dest: cbox) then have 1: "\\ \ lmeasurable" by blast have norme: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x,K)\p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" by (auto simp: lmeasure_integral_UNIV assms inteq dest: \) have "\x K y L. (x,K) \ (\K. (t K,K)) ` \ \ (y,L) \ (\K. (t K,K)) ` \ \ (x,K) \ (y,L) \ interior K \ interior L = {}" using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) with that \ have tagged: "(\K. (t K, K)) ` \ tagged_partial_division_of cbox a b" by (auto simp: tagged_partial_division_of_def dest: t cbox) have fine: "\ fine (\K. (t K, K)) ` \" using t by (auto simp: fine_def) have *: "y \ ?\ S \ \x - y\ \ e \ x \ ?\ S + e" for x y by arith have "?\ (\\) \ ?\ S + e" proof (rule *) have "(\K\\. ?\ (K \ S)) = ?\ (\C\\. C \ S)" apply (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) using \div \S \ lmeasurable\ apply blast unfolding pairwise_def by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) also have "\ = ?\ (\\ \ S)" by simp also have "\ \ ?\ S" by (simp add: "1" \S \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Int) finally show "(\K\\. ?\ (K \ S)) \ ?\ S" . next have "?\ (\\) = sum ?\ \" by (metis \div content_division) also have "\ = (\K\\. content K)" using \div by (force intro: sum.cong) also have "\ = (\x\\. content x * indicator S (t x))" using t by auto finally have eq1: "?\ (\\) = (\x\\. content x * indicator S (t x))" . have eq2: "(\K\\. ?\ (K \ S)) = (\K\\. integral K (indicator S))" apply (rule sum.cong [OF refl]) by (metis integral_indicator \div \S \ lmeasurable\ division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) have "\\(x,K)\(\K. (t K, K)) ` \. content K * indicator S x - integral K (indicator S)\ \ e" using Henstock_lemma_part1 [of "indicator S::'a\real", OF _ \e > 0\ \gauge \\ _ tagged fine] indS_int norme by auto then show "\?\ (\\) - (\K\\. ?\ (K \ S))\ \ e" by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) qed with 1 show ?thesis by blast qed have "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" proof (cases "finite \") case True with fincase show ?thesis by blast next case False let ?T = "from_nat_into \" have T: "bij_betw ?T UNIV \" by (simp add: False \(1) bij_betw_from_nat_into) have TM: "\n. ?T n \ lmeasurable" by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) have TN: "\m n. m \ n \ negligible (?T m \ ?T n)" by (simp add: False \(1) from_nat_into infinite_imp_nonempty negl_int) have TB: "(\k\n. ?\ (?T k)) \ ?\ S + e" for n proof - have "(\k\n. ?\ (?T k)) = ?\ (\ (?T ` {..n}))" by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) also have "?\ (\ (?T ` {..n})) \ ?\ S + e" using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def) finally show ?thesis . qed have "\\ \ lmeasurable" by (metis lmeasurable_compact T \(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) moreover have "?\ (\x. from_nat_into \ x) \ ?\ S + e" proof (rule measure_countable_Union_le [OF TM]) show "?\ (\x\n. from_nat_into \ x) \ ?\ S + e" for n by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) qed ultimately show ?thesis by (metis T bij_betw_def) qed then show "\\ \ lmeasurable" "measure lebesgue (\\) \ ?\ S + e" by blast+ qed (use \ cbox djointish close covers in auto) qed subsection\Transformation of measure by linear maps\ lemma emeasure_lebesgue_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "emeasure lebesgue (ball c r) = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" proof (cases "r = 0") case False with assms have r: "r > 0" by auto have "emeasure lebesgue ((\x. c + x) ` (\x. r *\<^sub>R x) ` ball (0 :: 'a) 1) = r ^ DIM('a) * emeasure lebesgue (ball (0 :: 'a) 1)" unfolding image_image using emeasure_lebesgue_affine[of r c "ball 0 1"] assms by (simp add: add_ac) also have "(\x. r *\<^sub>R x) ` ball 0 1 = ball (0 :: 'a) r" using r by (subst ball_scale) auto also have "(\x. c + x) ` \ = ball c r" by (subst image_add_ball) (simp_all add: algebra_simps) finally show ?thesis by simp qed auto lemma content_ball_conv_unit_ball: fixes c :: "'a :: euclidean_space" assumes "r \ 0" shows "content (ball c r) = r ^ DIM('a) * content (ball (0 :: 'a) 1)" proof - have "ennreal (content (ball c r)) = emeasure lebesgue (ball c r)" using emeasure_lborel_ball_finite[of c r] by (subst emeasure_eq_ennreal_measure) auto also have "\ = ennreal (r ^ DIM('a)) * emeasure lebesgue (ball (0 :: 'a) 1)" using assms by (intro emeasure_lebesgue_ball_conv_unit_ball) auto also have "\ = ennreal (r ^ DIM('a) * content (ball (0::'a) 1))" using emeasure_lborel_ball_finite[of "0::'a" 1] assms by (subst emeasure_eq_ennreal_measure) (auto simp: ennreal_mult') finally show ?thesis using assms by (subst (asm) ennreal_inj) auto qed lemma measurable_linear_image_interval: "linear f \ f ` (cbox a b) \ lmeasurable" by (metis bounded_linear_image linear_linear bounded_cbox closure_bounded_linear_image closure_cbox compact_closure lmeasurable_compact) proposition measure_linear_sufficient: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" and S: "S \ lmeasurable" and im: "\a b. measure lebesgue (f ` (cbox a b)) = m * measure lebesgue (cbox a b)" shows "f ` S \ lmeasurable \ m * measure lebesgue S = measure lebesgue (f ` S)" using le_less_linear [of 0 m] proof assume "m < 0" then show ?thesis using im [of 0 One] by auto next assume "m \ 0" let ?\ = "measure lebesgue" show ?thesis proof (cases "inj f") case False then have "?\ (f ` S) = 0" using \linear f\ negligible_imp_measure0 negligible_linear_singular_image by blast then have "m * ?\ (cbox 0 (One)) = 0" by (metis False \linear f\ cbox_borel content_unit im measure_completion negligible_imp_measure0 negligible_linear_singular_image sets_lborel) then show ?thesis using \linear f\ negligible_linear_singular_image negligible_imp_measure0 False by (auto simp: lmeasurable_iff_has_integral negligible_UNIV) next case True then obtain h where "linear h" and hf: "\x. h (f x) = x" and fh: "\x. f (h x) = x" using \linear f\ linear_injective_isomorphism by blast have fBS: "(f ` S) \ lmeasurable \ m * ?\ S = ?\ (f ` S)" if "bounded S" "S \ lmeasurable" for S proof - obtain a b where "S \ cbox a b" using \bounded S\ bounded_subset_cbox_symmetric by metis have fUD: "(f ` \\) \ lmeasurable \ ?\ (f ` \\) = (m * ?\ (\\))" if "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intint: "pairwise (\A B. interior A \ interior B = {}) \" for \ proof - have conv: "\K. K \ \ \ convex K" using cbox convex_box(1) by blast have neg: "negligible (g ` K \ g ` L)" if "linear g" "K \ \" "L \ \" "K \ L" for K L and g :: "'n\'n" proof (cases "inj g") case True have "negligible (frontier(g ` K \ g ` L) \ interior(g ` K \ g ` L))" proof (rule negligible_Un) show "negligible (frontier (g ` K \ g ` L))" by (simp add: negligible_convex_frontier convex_Int conv convex_linear_image that) next have "\p N. pairwise p N = (\Na. (Na::'n set) \ N \ (\Nb. Nb \ N \ Na \ Nb \ p Na Nb))" by (metis pairwise_def) then have "interior K \ interior L = {}" using intint that(2) that(3) that(4) by presburger then show "negligible (interior (g ` K \ g ` L))" by (metis True empty_imp_negligible image_Int image_empty interior_Int interior_injective_linear_image that(1)) qed moreover have "g ` K \ g ` L \ frontier (g ` K \ g ` L) \ interior (g ` K \ g ` L)" apply (auto simp: frontier_def) using closure_subset contra_subsetD by fastforce+ ultimately show ?thesis by (rule negligible_subset) next case False then show ?thesis by (simp add: negligible_Int negligible_linear_singular_image \linear g\) qed have negf: "negligible ((f ` K) \ (f ` L))" and negid: "negligible (K \ L)" if "K \ \" "L \ \" "K \ L" for K L using neg [OF \linear f\] neg [OF linear_id] that by auto show ?thesis proof (cases "finite \") case True then have "?\ (\x\\. f ` x) = (\x\\. ?\ (f ` x))" using \linear f\ cbox measurable_linear_image_interval negf by (blast intro: measure_negligible_finite_Union_image [unfolded pairwise_def]) also have "\ = (\k\\. m * ?\ k)" by (metis (no_types, lifting) cbox im sum.cong) also have "\ = m * ?\ (\\)" unfolding sum_distrib_left [symmetric] by (metis True cbox lmeasurable_cbox measure_negligible_finite_Union [unfolded pairwise_def] negid) finally show ?thesis by (metis True \linear f\ cbox image_Union fmeasurable.finite_UN measurable_linear_image_interval) next case False with \countable \\ obtain X :: "nat \ 'n set" where S: "bij_betw X UNIV \" using bij_betw_from_nat_into by blast then have eq: "(\\) = (\n. X n)" "(f ` \\) = (\n. f ` X n)" by (auto simp: bij_betw_def) have meas: "\K. K \ \ \ K \ lmeasurable" using cbox by blast with S have 1: "\n. X n \ lmeasurable" by (auto simp: bij_betw_def) have 2: "pairwise (\m n. negligible (X m \ X n)) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negid range_eqI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have 3: "bounded (\n. X n)" using S unfolding bij_betw_def by blast have "(\n. X n) \ lmeasurable" by (rule measurable_countable_negligible_Union_bounded [OF 1 2 3]) with S have f1: "\n. f ` (X n) \ lmeasurable" unfolding bij_betw_def by (metis assms(1) cbox measurable_linear_image_interval rangeI) have f2: "pairwise (\m n. negligible (f ` (X m) \ f ` (X n))) UNIV" using S unfolding bij_betw_def pairwise_def by (metis injD negf rangeI) have "bounded (\\)" by (meson Sup_least bounded_cbox bounded_subset cbox) then have f3: "bounded (\n. f ` X n)" using S unfolding bij_betw_def by (metis bounded_linear_image linear_linear assms(1) image_Union range_composition) have "(\n. ?\ (X n)) sums ?\ (\n. X n)" by (rule measure_countable_negligible_Union_bounded [OF 1 2 3]) have meq: "?\ (\n. f ` X n) = m * ?\ (\(X ` UNIV))" proof (rule sums_unique2 [OF measure_countable_negligible_Union_bounded [OF f1 f2 f3]]) have m: "\n. ?\ (f ` X n) = (m * ?\ (X n))" using S unfolding bij_betw_def by (metis cbox im rangeI) show "(\n. ?\ (f ` X n)) sums (m * ?\ (\(X ` UNIV)))" unfolding m using measure_countable_negligible_Union_bounded [OF 1 2 3] sums_mult by blast qed show ?thesis using measurable_countable_negligible_Union_bounded [OF f1 f2 f3] meq by (auto simp: eq [symmetric]) qed qed show ?thesis unfolding completion.fmeasurable_measure_inner_outer_le proof (intro conjI allI impI) fix e :: real assume "e > 0" have 1: "cbox a b - S \ lmeasurable" by (simp add: fmeasurable.Diff that) have 2: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "cbox a b - S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ (cbox a b - S) + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of "cbox a b - S" a b "e/(1 + \m\)"]; use 1 2 pairwise_def in force) have meq: "?\ (cbox a b - S) = ?\ (cbox a b) - ?\ S" by (simp add: measurable_measure_Diff \S \ cbox a b\ fmeasurableD that(2)) show "\T \ lmeasurable. T \ f ` S \ m * ?\ S - e \ ?\ T" proof (intro bexI conjI) show "f ` (cbox a b) - f ` (\\) \ f ` S" using \cbox a b - S \ \\\ by force have "m * ?\ S - e \ m * (?\ S - e / (1 + \m\))" using \m \ 0\ \e > 0\ by (simp add: field_simps) also have "\ \ ?\ (f ` cbox a b) - ?\ (f ` (\\))" using le \m \ 0\ \e > 0\ apply (simp add: im fUD [OF \countable \\ cbox intdisj] right_diff_distrib [symmetric]) apply (rule mult_left_mono; simp add: algebra_simps meq) done also have "\ = ?\ (f ` cbox a b - f ` \\)" apply (rule measurable_measure_Diff [symmetric]) apply (simp add: assms(1) measurable_linear_image_interval) apply (simp add: \countable \\ cbox fUD fmeasurableD intdisj) apply (simp add: Sup_le_iff cbox image_mono) done finally show "m * ?\ S - e \ ?\ (f ` cbox a b - f ` \\)" . show "f ` cbox a b - f ` \\ \ lmeasurable" by (simp add: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) qed next fix e :: real assume "e > 0" have em: "0 < e / (1 + \m\)" using \e > 0\ by (simp add: field_split_simps abs_add_one_gt_zero) obtain \ where "countable \" and cbox: "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" and intdisj: "pairwise (\A B. interior A \ interior B = {}) \" and DD: "S \ \\" "\\ \ lmeasurable" and le: "?\ (\\) \ ?\ S + e/(1 + \m\)" by (rule measurable_outer_intervals_bounded [of S a b "e/(1 + \m\)"]; use \S \ lmeasurable\ \S \ cbox a b\ em in force) show "\U \ lmeasurable. f ` S \ U \ ?\ U \ m * ?\ S + e" proof (intro bexI conjI) show "f ` S \ f ` (\\)" by (simp add: DD(1) image_mono) have "?\ (f ` \\) \ m * (?\ S + e / (1 + \m\))" using \m \ 0\ le mult_left_mono by (auto simp: fUD \countable \\ \linear f\ cbox fmeasurable.Diff intdisj measurable_linear_image_interval) also have "\ \ m * ?\ S + e" using \m \ 0\ \e > 0\ by (simp add: fUD [OF \countable \\ cbox intdisj] field_simps) finally show "?\ (f ` \\) \ m * ?\ S + e" . show "f ` \\ \ lmeasurable" by (simp add: \countable \\ cbox fUD intdisj) qed qed qed show ?thesis unfolding has_measure_limit_iff proof (intro allI impI) fix e :: real assume "e > 0" obtain B where "B > 0" and B: "\a b. ball 0 B \ cbox a b \ \?\ (S \ cbox a b) - ?\ S\ < e / (1 + \m\)" using has_measure_limit [OF S] \e > 0\ by (metis abs_add_one_gt_zero zero_less_divide_iff) obtain c d::'n where cd: "ball 0 B \ cbox c d" by (metis bounded_subset_cbox_symmetric bounded_ball) with B have less: "\?\ (S \ cbox c d) - ?\ S\ < e / (1 + \m\)" . obtain D where "D > 0" and D: "cbox c d \ ball 0 D" by (metis bounded_cbox bounded_subset_ballD) obtain C where "C > 0" and C: "\x. norm (f x) \ C * norm x" using linear_bounded_pos \linear f\ by blast have "f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" if "ball 0 (D*C) \ cbox a b" for a b proof - have "bounded (S \ h ` cbox a b)" by (simp add: bounded_linear_image linear_linear \linear h\ bounded_Int) moreover have Shab: "S \ h ` cbox a b \ lmeasurable" by (simp add: S \linear h\ fmeasurable.Int measurable_linear_image_interval) moreover have fim: "f ` (S \ h ` (cbox a b)) = (f ` S) \ cbox a b" by (auto simp: hf rev_image_eqI fh) ultimately have 1: "(f ` S) \ cbox a b \ lmeasurable" and 2: "m * ?\ (S \ h ` cbox a b) = ?\ ((f ` S) \ cbox a b)" using fBS [of "S \ (h ` (cbox a b))"] by auto have *: "\\z - m\ < e; z \ w; w \ m\ \ \w - m\ \ e" for w z m and e::real by auto have meas_adiff: "\?\ (S \ h ` cbox a b) - ?\ S\ \ e / (1 + \m\)" proof (rule * [OF less]) show "?\ (S \ cbox c d) \ ?\ (S \ h ` cbox a b)" proof (rule measure_mono_fmeasurable [OF _ _ Shab]) have "f ` ball 0 D \ ball 0 (C * D)" using C \C > 0\ apply (clarsimp simp: algebra_simps) by (meson le_less_trans linordered_comm_semiring_strict_class.comm_mult_strict_left_mono) then have "f ` ball 0 D \ cbox a b" by (metis mult.commute order_trans that) have "ball 0 D \ h ` cbox a b" by (metis \f ` ball 0 D \ cbox a b\ hf image_subset_iff subsetI) then show "S \ cbox c d \ S \ h ` cbox a b" using D by blast next show "S \ cbox c d \ sets lebesgue" using S fmeasurable_cbox by blast qed next show "?\ (S \ h ` cbox a b) \ ?\ S" by (simp add: S Shab fmeasurableD measure_mono_fmeasurable) qed have "\?\ (f ` S \ cbox a b) - m * ?\ S\ \ m * e / (1 + \m\)" proof - have mm: "\m\ = m" by (simp add: \0 \ m\) then have "\?\ S - ?\ (S \ h ` cbox a b)\ * m \ e / (1 + m) * m" by (metis (no_types) \0 \ m\ meas_adiff abs_minus_commute mult_right_mono) moreover have "\r. \r * m\ = m * \r\" by (metis \0 \ m\ abs_mult_pos mult.commute) ultimately show ?thesis apply (simp add: 2 [symmetric]) by (metis (no_types) abs_minus_commute mult.commute right_diff_distrib' mm) qed also have "\ < e" using \e > 0\ by (cases "m \ 0") (simp_all add: field_simps) finally have "\?\ (f ` S \ cbox a b) - m * ?\ S\ < e" . with 1 show ?thesis by auto qed then show "\B>0. \a b. ball 0 B \ cbox a b \ f ` S \ cbox a b \ lmeasurable \ \?\ (f ` S \ cbox a b) - m * ?\ S\ < e" using \C>0\ \D>0\ by (metis mult_zero_left real_mult_less_iff1) qed qed qed subsection\Lemmas about absolute integrability\ lemma absolutely_integrable_linear: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and h :: "'n::euclidean_space \ 'p::euclidean_space" shows "f absolutely_integrable_on s \ bounded_linear h \ (h \ f) absolutely_integrable_on s" using integrable_bounded_linear[of h lebesgue "\x. indicator s x *\<^sub>R f x"] by (simp add: linear_simps[of h] set_integrable_def) lemma absolutely_integrable_sum: fixes f :: "'a \ 'n::euclidean_space \ 'm::euclidean_space" assumes "finite T" and "\a. a \ T \ (f a) absolutely_integrable_on S" shows "(\x. sum (\a. f a x) T) absolutely_integrable_on S" using assms by induction auto lemma absolutely_integrable_integrable_bound: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes le: "\x. x\S \ norm (f x) \ g x" and f: "f integrable_on S" and g: "g integrable_on S" shows "f absolutely_integrable_on S" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) have "g absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (g x)) integrable_on S" using le norm_ge_zero[of "f _"] by (intro integrable_spike_finite[OF _ _ g, of "{}"]) (auto intro!: abs_of_nonneg intro: order_trans simp del: norm_ge_zero) qed fact then show "integrable lebesgue (\x. indicat_real S x *\<^sub>R g x)" by (simp add: set_integrable_def) show "(\x. indicat_real S x *\<^sub>R f x) \ borel_measurable lebesgue" using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) qed (use le in \force intro!: always_eventually split: split_indicator\) corollary absolutely_integrable_on_const [simp]: fixes c :: "'a::euclidean_space" assumes "S \ lmeasurable" shows "(\x. c) absolutely_integrable_on S" by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) lemma absolutely_integrable_continuous: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "continuous_on (cbox a b) f \ f absolutely_integrable_on cbox a b" using absolutely_integrable_integrable_bound by (simp add: absolutely_integrable_on_def continuous_on_norm integrable_continuous) lemma absolutely_integrable_continuous_real: fixes f :: "real \ 'b::euclidean_space" shows "continuous_on {a..b} f \ f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_continuous box_real(2)) lemma continuous_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on (cbox a b) f" shows "integrable (lebesgue_on (cbox a b)) f" proof - have "f absolutely_integrable_on cbox a b" by (simp add: absolutely_integrable_continuous assms) then show ?thesis by (simp add: integrable_restrict_space set_integrable_def) qed lemma continuous_imp_integrable_real: fixes f :: "real \ 'b::euclidean_space" assumes "continuous_on {a..b} f" shows "integrable (lebesgue_on {a..b}) f" by (metis assms continuous_imp_integrable interval_cbox) subsection \Componentwise\ proposition absolutely_integrable_componentwise_iff: shows "f absolutely_integrable_on A \ (\b\Basis. (\x. f x \ b) absolutely_integrable_on A)" proof - have *: "(\x. norm (f x)) integrable_on A \ (\b\Basis. (\x. norm (f x \ b)) integrable_on A)" if "f integrable_on A" proof - have 1: "\i. \(\x. norm (f x)) integrable_on A; i \ Basis\ \ (\x. f x \ i) absolutely_integrable_on A" apply (rule absolutely_integrable_integrable_bound [where g = "\x. norm(f x)"]) using Basis_le_norm integrable_component that apply fastforce+ done have 2: "\i\Basis. (\x. \f x \ i\) integrable_on A \ f absolutely_integrable_on A" apply (rule absolutely_integrable_integrable_bound [where g = "\x. \i\Basis. norm (f x \ i)"]) using norm_le_l1 that apply (force intro: integrable_sum)+ done show ?thesis apply auto apply (metis (full_types) absolutely_integrable_on_def set_integrable_abs 1) apply (metis (full_types) absolutely_integrable_on_def 2) done qed show ?thesis unfolding absolutely_integrable_on_def by (simp add: integrable_componentwise_iff [symmetric] ball_conj_distrib * cong: conj_cong) qed lemma absolutely_integrable_componentwise: shows "(\b. b \ Basis \ (\x. f x \ b) absolutely_integrable_on A) \ f absolutely_integrable_on A" using absolutely_integrable_componentwise_iff by blast lemma absolutely_integrable_component: "f absolutely_integrable_on A \ (\x. f x \ (b :: 'b :: euclidean_space)) absolutely_integrable_on A" by (drule absolutely_integrable_linear[OF _ bounded_linear_inner_left[of b]]) (simp add: o_def) lemma absolutely_integrable_scaleR_left: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. c *\<^sub>R f x) absolutely_integrable_on S" proof - have "(\x. c *\<^sub>R x) o f absolutely_integrable_on S" apply (rule absolutely_integrable_linear [OF assms]) by (simp add: bounded_linear_scaleR_right) then show ?thesis using assms by blast qed lemma absolutely_integrable_scaleR_right: assumes "f absolutely_integrable_on S" shows "(\x. f x *\<^sub>R c) absolutely_integrable_on S" using assms by blast lemma absolutely_integrable_norm: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(norm o f) absolutely_integrable_on S" using assms by (simp add: absolutely_integrable_on_def o_def) lemma absolutely_integrable_abs: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f absolutely_integrable_on S" shows "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) absolutely_integrable_on S" (is "?g absolutely_integrable_on S") proof - have eq: "?g = (\x. \i\Basis. ((\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm(\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f) x)" by (simp add: sum.delta) have *: "(\y. \j\Basis. if j = i then y *\<^sub>R j else 0) \ (\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" if "i \ Basis" for i proof - have "bounded_linear (\y. \j\Basis. if j = i then y *\<^sub>R j else 0)" by (simp add: linear_linear algebra_simps linearI) moreover have "(\x. norm (\j\Basis. if j = i then (x \ i) *\<^sub>R j else 0)) \ f absolutely_integrable_on S" unfolding o_def apply (rule absolutely_integrable_norm [unfolded o_def]) using assms \i \ Basis\ apply (auto simp: algebra_simps dest: absolutely_integrable_component[where b=i]) done ultimately show ?thesis by (subst comp_assoc) (blast intro: absolutely_integrable_linear) qed show ?thesis apply (rule ssubst [OF eq]) apply (rule absolutely_integrable_sum) apply (force simp: intro!: *)+ done qed lemma abs_absolutely_integrableI_1: fixes f :: "'a :: euclidean_space \ real" assumes f: "f integrable_on A" and "(\x. \f x\) integrable_on A" shows "f absolutely_integrable_on A" by (rule absolutely_integrable_integrable_bound [OF _ assms]) auto lemma abs_absolutely_integrableI: assumes f: "f integrable_on S" and fcomp: "(\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" shows "f absolutely_integrable_on S" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" if "i \ Basis" for i proof - have "(\x. \f x \ i\) integrable_on S" using assms integrable_component [OF fcomp, where y=i] that by simp then have "(\x. f x \ i) absolutely_integrable_on S" using abs_absolutely_integrableI_1 f integrable_component by blast then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on S" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_abs_iff: "f absolutely_integrable_on S \ f integrable_on S \ (\x. \i\Basis. \f x \ i\ *\<^sub>R i) integrable_on S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using absolutely_integrable_abs absolutely_integrable_on_def by blast next assume ?rhs moreover have "(\x. if x \ S then \i\Basis. \f x \ i\ *\<^sub>R i else 0) = (\x. \i\Basis. \(if x \ S then f x else 0) \ i\ *\<^sub>R i)" by force ultimately show ?lhs by (simp only: absolutely_integrable_restrict_UNIV [of S, symmetric] integrable_restrict_UNIV [of S, symmetric] abs_absolutely_integrableI) qed lemma absolutely_integrable_max: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i + \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i + \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib algebra_simps euclidean_representation) finally show "(\i\Basis. max (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x + (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" apply (intro set_integral_add absolutely_integrable_scaleR_left assms) using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] apply (simp add: algebra_simps) done ultimately show ?thesis by metis qed corollary absolutely_integrable_max_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. max (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_max [OF assms] by simp lemma absolutely_integrable_min: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) absolutely_integrable_on S" proof - have "(\x. \i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\x. (1/2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i)))" proof (rule ext) fix x have "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (\i\Basis. ((f x \ i + g x \ i - \f x \ i - g x \ i\) / 2) *\<^sub>R i)" by (force intro: sum.cong) also have "... = (1 / 2) *\<^sub>R (\i\Basis. (f x \ i + g x \ i - \f x \ i - g x \ i\) *\<^sub>R i)" by (simp add: scaleR_right.sum) also have "... = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" by (simp add: sum.distrib sum_subtractf algebra_simps euclidean_representation) finally show "(\i\Basis. min (f x \ i) (g x \ i) *\<^sub>R i) = (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))" . qed moreover have "(\x. (1 / 2) *\<^sub>R (f x + g x - (\i\Basis. \f x \ i - g x \ i\ *\<^sub>R i))) absolutely_integrable_on S" apply (intro set_integral_add set_integral_diff absolutely_integrable_scaleR_left assms) using absolutely_integrable_abs [OF set_integral_diff(1) [OF assms]] apply (simp add: algebra_simps) done ultimately show ?thesis by metis qed corollary absolutely_integrable_min_1: fixes f :: "'n::euclidean_space \ real" assumes "f absolutely_integrable_on S" "g absolutely_integrable_on S" shows "(\x. min (f x) (g x)) absolutely_integrable_on S" using absolutely_integrable_min [OF assms] by simp lemma nonnegative_absolutely_integrable: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ 0 \ f x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" if "i \ Basis" for i proof - have "(\x. f x \ i) integrable_on A" by (simp add: assms(1) integrable_component) then have "(\x. f x \ i) absolutely_integrable_on A" by (metis that comp nonnegative_absolutely_integrable_1) then show ?thesis by (rule absolutely_integrable_scaleR_right) qed then have "(\x. \i\Basis. (f x \ i) *\<^sub>R i) absolutely_integrable_on A" by (simp add: absolutely_integrable_sum) then show ?thesis by (simp add: euclidean_representation) qed lemma absolutely_integrable_component_ubound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f integrable_on A" and g: "g absolutely_integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "f absolutely_integrable_on A" proof - have "(\x. g x - (g x - f x)) absolutely_integrable_on A" apply (rule set_integral_diff [OF g nonnegative_absolutely_integrable]) using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast by (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma absolutely_integrable_component_lbound: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f absolutely_integrable_on A" and g: "g integrable_on A" and comp: "\x b. \x \ A; b \ Basis\ \ f x \ b \ g x \ b" shows "g absolutely_integrable_on A" proof - have "(\x. f x + (g x - f x)) absolutely_integrable_on A" apply (rule set_integral_add [OF f nonnegative_absolutely_integrable]) using Henstock_Kurzweil_Integration.integrable_diff absolutely_integrable_on_def f g apply blast by (simp add: comp inner_diff_left) then show ?thesis by simp qed lemma integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f integrable_on S \ (\x. f x $ 1) integrable_on S" by (auto simp: integrable_componentwise_iff [of f] Basis_vec_def cart_eq_inner_axis) lemma integral_on_1_eq: fixes f :: "'a::euclidean_space \ real^1" shows "integral S f = vec (integral S (\x. f x $ 1))" by (cases "f integrable_on S") (simp_all add: integrable_on_1_iff vec_eq_iff not_integrable_integral) lemma absolutely_integrable_on_1_iff: fixes f :: "'a::euclidean_space \ real^1" shows "f absolutely_integrable_on S \ (\x. f x $ 1) absolutely_integrable_on S" unfolding absolutely_integrable_on_def by (auto simp: integrable_on_1_iff norm_real) lemma absolutely_integrable_absolutely_integrable_lbound: fixes f :: "'m::euclidean_space \ real" assumes f: "f integrable_on S" and g: "g absolutely_integrable_on S" and *: "\x. x \ S \ g x \ f x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_lbound [OF g f]) (simp add: *) lemma absolutely_integrable_absolutely_integrable_ubound: fixes f :: "'m::euclidean_space \ real" assumes fg: "f integrable_on S" "g absolutely_integrable_on S" and *: "\x. x \ S \ f x \ g x" shows "f absolutely_integrable_on S" by (rule absolutely_integrable_component_ubound [OF fg]) (simp add: *) lemma has_integral_vec1_I_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) (cbox a b)" shows "((f \ vec) has_integral y) {a$1..b$1}" proof - have "((\x. f(vec x)) has_integral (1 / 1) *\<^sub>R y) ((\x. x $ 1) ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real^1. vec ` cbox u v = cbox w z" "content (vec ` cbox u v :: (real^1) set) = 1 * content (cbox u v)" for u v unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z. (\x. x $ 1) ` cbox u v = cbox w z" for u v :: "real^1" using vec_nth_cbox_1_eq by blast qed (auto simp: continuous_vec assms) then show ?thesis by (simp add: o_def) qed lemma has_integral_vec1_I: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "(f has_integral y) S" shows "(f \ vec has_integral y) ((\x. x $ 1) ` S)" proof - have *: "\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e" if int: "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)" and B: "ball 0 B \ {a..b}" for e B a b proof - have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force have B': "ball (0::real^1) B \ cbox (vec a) (vec b)" using B by (simp add: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box norm_real subset_iff) show ?thesis using int [OF B'] by (auto simp: image_iff o_def cong: if_cong dest!: has_integral_vec1_I_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_I_cbox split: if_split_asm) apply (metis vector_one_nth) apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+ apply (blast intro!: *) done qed lemma has_integral_vec1_nth_cbox: fixes f :: "real \ 'a::real_normed_vector" assumes "(f has_integral y) {a..b}" shows "((\x::real^1. f(x$1)) has_integral y) (cbox (vec a) (vec b))" proof - have "((\x::real^1. f(x$1)) has_integral (1 / 1) *\<^sub>R y) (vec ` cbox a b)" proof (rule has_integral_twiddle) show "\w z::real. (\x. x $ 1) ` cbox u v = cbox w z" "content ((\x. x $ 1) ` cbox u v) = 1 * content (cbox u v)" for u v::"real^1" unfolding vec_cbox_1_eq by (auto simp: content_cbox_if_cart interval_eq_empty_cart) show "\w z::real^1. vec ` cbox u v = cbox w z" for u v :: "real" using vec_cbox_1_eq by auto qed (auto simp: continuous_vec assms) then show ?thesis using vec_cbox_1_eq by auto qed lemma has_integral_vec1_D_cbox: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) {a$1..b$1}" shows "(f has_integral y) (cbox a b)" by (metis (mono_tags, lifting) assms comp_apply has_integral_eq has_integral_vec1_nth_cbox vector_one_nth) lemma has_integral_vec1_D: fixes f :: "real^1 \ 'a::real_normed_vector" assumes "((f \ vec) has_integral y) ((\x. x $ 1) ` S)" shows "(f has_integral y) S" proof - have *: "\z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e" if int: "\a b. ball 0 B \ {a..b} \ (\z. ((\x. if x \ (\x. x $ 1) ` S then (f \ vec) x else 0) has_integral z) {a..b} \ norm (z - y) < e)" and B: "ball 0 B \ cbox a b" for e B and a b::"real^1" proof - have B': "ball 0 B \ {a$1..b$1}" using B apply (simp add: subset_iff Basis_vec_def cart_eq_inner_axis [symmetric] mem_box) apply (metis (full_types) norm_real vec_component) done have eq: "(\x. if vec x \ S then f (vec x) else 0) = (\x. if x \ S then f x else 0) \ vec" by force have [simp]: "(\y\S. x = y $ 1) \ vec x \ S" for x by force show ?thesis using int [OF B'] by (auto simp: image_iff eq cong: if_cong dest!: has_integral_vec1_D_cbox) qed show ?thesis using assms apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp add: has_integral_vec1_D_cbox eq_cbox split: if_split_asm, blast) apply (intro conjI impI) apply (metis vector_one_nth) apply (erule thin_rl) apply (erule all_forward imp_forward asm_rl ex_forward conj_forward)+ using * apply blast done qed lemma integral_vec1_eq: fixes f :: "real^1 \ 'a::real_normed_vector" shows "integral S f = integral ((\x. x $ 1) ` S) (f \ vec)" using has_integral_vec1_I [of f] has_integral_vec1_D [of f] by (metis has_integral_iff not_integrable_integral) lemma absolutely_integrable_drop: fixes f :: "real^1 \ 'b::euclidean_space" shows "f absolutely_integrable_on S \ (f \ vec) absolutely_integrable_on (\x. x $ 1) ` S" unfolding absolutely_integrable_on_def integrable_on_def proof safe fix y r assume "(f has_integral y) S" "((\x. norm (f x)) has_integral r) S" then show "\y. (f \ vec has_integral y) ((\x. x $ 1) ` S)" "\y. ((\x. norm ((f \ vec) x)) has_integral y) ((\x. x $ 1) ` S)" by (force simp: o_def dest!: has_integral_vec1_I)+ next fix y :: "'b" and r :: "real" assume "(f \ vec has_integral y) ((\x. x $ 1) ` S)" "((\x. norm ((f \ vec) x)) has_integral r) ((\x. x $ 1) ` S)" then show "\y. (f has_integral y) S" "\y. ((\x. norm (f x)) has_integral y) S" by (force simp: o_def intro: has_integral_vec1_D)+ qed subsection \Dominated convergence\ lemma dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. (f k) integrable_on S" and h: "h integrable_on S" and le: "\k x. x \ S \ norm (f k x) \ h x" and conv: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" "(\k. integral S (f k)) \ integral S g" proof - have 3: "h absolutely_integrable_on S" unfolding absolutely_integrable_on_def proof show "(\x. norm (h x)) integrable_on S" proof (intro integrable_spike_finite[OF _ _ h, of "{}"] ballI) fix x assume "x \ S - {}" then show "norm (h x) = h x" by (metis Diff_empty abs_of_nonneg bot_set_def le norm_ge_zero order_trans real_norm_def) qed auto qed fact have 2: "set_borel_measurable lebesgue S (f k)" for k unfolding set_borel_measurable_def using f by (auto intro: has_integral_implies_lebesgue_measurable simp: integrable_on_def) then have 1: "set_borel_measurable lebesgue S g" unfolding set_borel_measurable_def by (rule borel_measurable_LIMSEQ_metric) (use conv in \auto split: split_indicator\) have 4: "AE x in lebesgue. (\i. indicator S x *\<^sub>R f i x) \ indicator S x *\<^sub>R g x" "AE x in lebesgue. norm (indicator S x *\<^sub>R f k x) \ indicator S x *\<^sub>R h x" for k using conv le by (auto intro!: always_eventually split: split_indicator) have g: "g absolutely_integrable_on S" using 1 2 3 4 unfolding set_borel_measurable_def set_integrable_def by (rule integrable_dominated_convergence) then show "g integrable_on S" by (auto simp: absolutely_integrable_on_def) have "(\k. (LINT x:S|lebesgue. f k x)) \ (LINT x:S|lebesgue. g x)" unfolding set_borel_measurable_def set_lebesgue_integral_def using 1 2 3 4 unfolding set_borel_measurable_def set_lebesgue_integral_def set_integrable_def by (rule integral_dominated_convergence) then show "(\k. integral S (f k)) \ integral S g" using g absolutely_integrable_integrable_bound[OF le f h] by (subst (asm) (1 2) set_lebesgue_integral_eq_integral) auto qed lemma has_integral_dominated_convergence: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes "\k. (f k has_integral y k) S" "h integrable_on S" "\k. \x\S. norm (f k x) \ h x" "\x\S. (\k. f k x) \ g x" and x: "y \ x" shows "(g has_integral x) S" proof - have int_f: "\k. (f k) integrable_on S" using assms by (auto simp: integrable_on_def) have "(g has_integral (integral S g)) S" by (metis assms(2-4) dominated_convergence(1) has_integral_integral int_f) moreover have "integral S g = x" proof (rule LIMSEQ_unique) show "(\i. integral S (f i)) \ x" using integral_unique[OF assms(1)] x by simp show "(\i. integral S (f i)) \ integral S g" by (metis assms(2) assms(3) assms(4) dominated_convergence(2) int_f) qed ultimately show ?thesis by simp qed lemma dominated_convergence_integrable_1: fixes f :: "nat \ 'n::euclidean_space \ real" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" proof - have habs: "h absolutely_integrable_on S" using h normg nonnegative_absolutely_integrable_1 norm_ge_zero order_trans by blast let ?f = "\n x. (min (max (- h x) (f n x)) (h x))" have h0: "h x \ 0" if "x \ S" for x using normg that by force have leh: "norm (?f k x) \ h x" if "x \ S" for k x using h0 that by force have limf: "(\k. ?f k x) \ g x" if "x \ S" for x proof - have "\e y. \f y x - g x\ < e \ \min (max (- h x) (f y x)) (h x) - g x\ < e" using h0 [OF that] normg [OF that] by simp then show ?thesis using lim [OF that] by (auto simp add: tendsto_iff dist_norm elim!: eventually_mono) qed show ?thesis proof (rule dominated_convergence [of ?f S h g]) have "(\x. - h x) absolutely_integrable_on S" using habs unfolding set_integrable_def by auto then show "?f k integrable_on S" for k by (intro set_lebesgue_integral_eq_integral absolutely_integrable_min_1 absolutely_integrable_max_1 f habs) qed (use assms leh limf in auto) qed lemma dominated_convergence_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g integrable_on S" using f unfolding integrable_componentwise_iff [of g] absolutely_integrable_componentwise_iff [where f = "f k" for k] proof clarify fix b :: "'m" assume fb [rule_format]: "\k. \b\Basis. (\x. f k x \ b) absolutely_integrable_on S" and b: "b \ Basis" show "(\x. g x \ b) integrable_on S" proof (rule dominated_convergence_integrable_1 [OF fb h]) fix x assume "x \ S" show "norm (g x \ b) \ h x" using norm_nth_le \x \ S\ b normg order.trans by blast show "(\k. f k x \ b) \ g x \ b" using \x \ S\ b lim tendsto_componentwise_iff by fastforce qed (use b in auto) qed lemma dominated_convergence_absolutely_integrable: fixes f :: "nat \ 'n::euclidean_space \ 'm::euclidean_space" assumes f: "\k. f k absolutely_integrable_on S" and h: "h integrable_on S" and normg: "\x. x \ S \ norm(g x) \ (h x)" and lim: "\x. x \ S \ (\k. f k x) \ g x" shows "g absolutely_integrable_on S" proof - have "g integrable_on S" by (rule dominated_convergence_integrable [OF assms]) with assms show ?thesis by (blast intro: absolutely_integrable_integrable_bound [where g=h]) qed proposition integral_countable_UN: fixes f :: "real^'m \ real^'n" assumes f: "f absolutely_integrable_on (\(range s))" and s: "\m. s m \ sets lebesgue" shows "\n. f absolutely_integrable_on (\m\n. s m)" and "(\n. integral (\m\n. s m) f) \ integral (\(s ` UNIV)) f" (is "?F \ ?I") proof - show fU: "f absolutely_integrable_on (\m\n. s m)" for n using assms by (blast intro: set_integrable_subset [OF f]) have fint: "f integrable_on (\ (range s))" using absolutely_integrable_on_def f by blast let ?h = "\x. if x \ \(s ` UNIV) then norm(f x) else 0" have "(\n. integral UNIV (\x. if x \ (\m\n. s m) then f x else 0)) \ integral UNIV (\x. if x \ \(s ` UNIV) then f x else 0)" proof (rule dominated_convergence) show "(\x. if x \ (\m\n. s m) then f x else 0) integrable_on UNIV" for n unfolding integrable_restrict_UNIV using fU absolutely_integrable_on_def by blast show "(\x. if x \ \(s ` UNIV) then norm(f x) else 0) integrable_on UNIV" by (metis (no_types) absolutely_integrable_on_def f integrable_restrict_UNIV) show "\x. (\n. if x \ (\m\n. s m) then f x else 0) \ (if x \ \(s ` UNIV) then f x else 0)" by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto then show "?F \ ?I" by (simp only: integral_restrict_UNIV) qed subsection \Fundamental Theorem of Calculus for the Lebesgue integral\ text \ For the positive integral we replace continuity with Borel-measurability. \ lemma fixes f :: "real \ real" assumes [measurable]: "f \ borel_measurable borel" assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" shows nn_integral_FTC_Icc: "(\\<^sup>+x. ennreal (f x) * indicator {a .. b} x \lborel) = F b - F a" (is ?nn) and has_bochner_integral_FTC_Icc_nonneg: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_nonneg: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) and integrable_FTC_Icc_nonneg: "integrable lborel (\x. f x * indicator {a .. b} x)" (is ?int) proof - have *: "(\x. f x * indicator {a..b} x) \ borel_measurable borel" "\x. 0 \ f x * indicator {a..b} x" using f(2) by (auto split: split_indicator) have F_mono: "a \ x \ x \ y \ y \ b\ F x \ F y" for x y using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) have "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] intro: has_field_derivative_subset[OF f(1)] \a \ b\) then have i: "((\x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV" unfolding indicator_def if_distrib[where f="\x. a * x" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) then have nn: "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" by (rule nn_integral_has_integral_lborel[OF *]) then show ?has by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \a \ b\) then show ?eq ?int unfolding has_bochner_integral_iff by auto show ?nn by (subst nn[symmetric]) (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator) qed lemma fixes f :: "real \ 'a :: euclidean_space" assumes "a \ b" assumes "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" assumes cont: "continuous_on {a .. b} f" shows has_bochner_integral_FTC_Icc: "has_bochner_integral lborel (\x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has) and integral_FTC_Icc: "(\x. indicator {a .. b} x *\<^sub>R f x \lborel) = F b - F a" (is ?eq) proof - let ?f = "\x. indicator {a .. b} x *\<^sub>R f x" have int: "integrable lborel ?f" using borel_integrable_compact[OF _ cont] by auto have "(f has_integral F b - F a) {a..b}" using assms(1,2) by (intro fundamental_theorem_of_calculus) auto moreover have "(f has_integral integral\<^sup>L lborel ?f) {a..b}" using has_integral_integral_lborel[OF int] unfolding indicator_def if_distrib[where f="\x. x *\<^sub>R a" for a] by (simp cong del: if_weak_cong del: atLeastAtMost_iff) ultimately show ?eq by (auto dest: has_integral_unique) then show ?has using int by (auto simp: has_bochner_integral_iff) qed lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ DERIV F x :> f x" assumes cont: "\x. a \ x \ x \ b \ isCont f x" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_field_derivative_iff_has_vector_derivative[symmetric] using deriv by (auto intro: DERIV_subset) have 2: "continuous_on {a .. b} f" using cont by (intro continuous_at_imp_continuous_on) auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 2] integral_FTC_Icc[OF \a \ b\ 1 2] by (auto simp: mult.commute) qed lemma nn_integral_FTC_atLeast: fixes f :: "real \ real" assumes f_borel: "f \ borel_measurable borel" assumes f: "\x. a \ x \ DERIV F x :> f x" assumes nonneg: "\x. a \ x \ 0 \ f x" assumes lim: "(F \ T) at_top" shows "(\\<^sup>+x. ennreal (f x) * indicator {a ..} x \lborel) = T - F a" proof - let ?f = "\(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x" let ?fR = "\x. ennreal (f x) * indicator {a ..} x" have F_mono: "a \ x \ x \ y \ F x \ F y" for x y using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans) then have F_le_T: "a \ x \ F x \ T" for x by (intro tendsto_lowerbound[OF lim]) (auto simp: eventually_at_top_linorder) have "(SUP i. ?f i x) = ?fR x" for x proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) obtain n where "x - a < real n" using reals_Archimedean2[of "x - a"] .. then have "eventually (\n. ?f n x = ?fR x) sequentially" by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) then show "(\n. ?f n x) \ ?fR x" by (rule tendsto_eventually) qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator) then have "integral\<^sup>N lborel ?fR = (\\<^sup>+ x. (SUP i. ?f i x) \lborel)" by simp also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \lborel))" proof (rule nn_integral_monotone_convergence_SUP) show "incseq ?f" using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) show "\i. (?f i) \ borel_measurable lborel" using f_borel by auto qed also have "\ = (SUP i. ennreal (F (a + real i) - F a))" by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto also have "\ = T - F a" proof (rule LIMSEQ_unique[OF LIMSEQ_SUP]) have "(\x. F (a + real x)) \ T" apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl]) apply (rule filterlim_real_sequentially) done then show "(\n. ennreal (F (a + real n) - F a)) \ ennreal (T - F a)" by (simp add: F_mono F_le_T tendsto_diff) qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono) finally show ?thesis . qed lemma integral_power: "a \ b \ (\x. x^k * indicator {a..b} x \lborel) = (b^Suc k - a^Suc k) / Suc k" proof (subst integral_FTC_Icc_real) fix x show "DERIV (\x. x^Suc k / Suc k) x :> x^k" by (intro derivative_eq_intros) auto qed (auto simp: field_simps simp del: of_nat_Suc) subsection \Integration by parts\ lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "!!x. a \x \ x\b \ isCont f x" assumes cont_g[intro]: "!!x. a \x \ x\b \ isCont g x" assumes [intro]: "!!x. DERIV F x :> f x" assumes [intro]: "!!x. DERIV G x :> g x" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: DERIV_isCont) have "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = (\x. (F x * g x) * indicator {a .. b} x \lborel) + \x. (f x * G x) * indicator {a .. b} x \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros) by (auto intro!: DERIV_isCont Bochner_Integration.integral_cong split: split_indicator) thus ?thesis using 0 by auto qed lemma integral_by_parts': fixes f g F G::"real \ real" assumes "a \ b" assumes "!!x. a \x \ x\b \ isCont f x" assumes "!!x. a \x \ x\b \ isCont g x" assumes "!!x. DERIV F x :> f x" assumes "!!x. DERIV G x :> g x" shows "(\x. indicator {a .. b} x *\<^sub>R (F x * g x) \lborel) = F b * G b - F a * G a - \x. indicator {a .. b} x *\<^sub>R (f x * G x) \lborel" using integral_by_parts[OF assms] by (simp add: ac_simps) lemma has_bochner_integral_even_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes even: "\x. f (- x) = f x" shows "has_bochner_integral lborel f (2 *\<^sub>R x)" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator even f) with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by (simp add: scaleR_2) qed lemma has_bochner_integral_odd_function: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x) x" assumes odd: "\x. f (- x) = - f x" shows "has_bochner_integral lborel f 0" proof - have indicator: "\x::real. indicator {..0} (- x) = indicator {0..} x" by (auto split: split_indicator) have "has_bochner_integral lborel (\x. - indicator {.. 0} x *\<^sub>R f x) x" by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) (auto simp: indicator odd f) from has_bochner_integral_minus[OF this] have "has_bochner_integral lborel (\x. indicator {.. 0} x *\<^sub>R f x) (- x)" by simp with f have "has_bochner_integral lborel (\x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)" by (rule has_bochner_integral_add) then have "has_bochner_integral lborel f (x + - x)" by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) (auto split: split_indicator) then show ?thesis by simp qed lemma has_integral_0_closure_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (closure S) f" and nonneg_interior: "\x. x \ S \ 0 \ f x" and pos: "0 < emeasure lborel S" and finite: "emeasure lborel S < \" and regular: "emeasure lborel (closure S) = emeasure lborel S" and opn: "open S" assumes int: "(f has_integral 0) (closure S)" assumes x: "x \ closure S" shows "f x = 0" proof - have zero: "emeasure lborel (frontier S) = 0" using finite closure_subset regular unfolding frontier_def by (subst emeasure_Diff) (auto simp: frontier_def interior_open \open S\ ) have nonneg: "0 \ f x" if "x \ closure S" for x using continuous_ge_on_closure[OF f that nonneg_interior] by simp have "0 = integral (closure S) f" by (blast intro: int sym) also note intl = has_integral_integrable[OF int] have af: "f absolutely_integrable_on (closure S)" using nonneg by (intro absolutely_integrable_onI intl integrable_eq[OF intl]) simp then have "integral (closure S) f = set_lebesgue_integral lebesgue (closure S) f" by (intro set_lebesgue_integral_eq_integral(2)[symmetric]) also have "\ = 0 \ (AE x in lebesgue. indicator (closure S) x *\<^sub>R f x = 0)" unfolding set_lebesgue_integral_def proof (rule integral_nonneg_eq_0_iff_AE) show "integrable lebesgue (\x. indicat_real (closure S) x *\<^sub>R f x)" by (metis af set_integrable_def) qed (use nonneg in \auto simp: indicator_def\) also have "\ \ (AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by (auto simp: indicator_def) finally have "(AE x in lebesgue. x \ {x. x \ closure S \ f x = 0})" by simp moreover have "(AE x in lebesgue. x \ - frontier S)" using zero by (auto simp: eventually_ae_filter null_sets_def intro!: exI[where x="frontier S"]) ultimately have ae: "AE x \ S in lebesgue. x \ {x \ closure S. f x = 0}" (is ?th) by eventually_elim (use closure_subset in \auto simp: \) have "closed {0::real}" by simp with continuous_on_closed_vimage[OF closed_closure, of S f] f have "closed (f -` {0} \ closure S)" by blast then have "closed {x \ closure S. f x = 0}" by (auto simp: vimage_def Int_def conj_commute) with \open S\ have "x \ {x \ closure S. f x = 0}" if "x \ S" for x using ae that by (rule mem_closed_if_AE_lebesgue_open) then have "f x = 0" if "x \ S" for x using that by auto from continuous_constant_on_closure[OF f this \x \ closure S\] show "f x = 0" . qed lemma has_integral_0_cbox_imp_0: fixes f :: "'a::euclidean_space \ real" assumes f: "continuous_on (cbox a b) f" and nonneg_interior: "\x. x \ box a b \ 0 \ f x" assumes int: "(f has_integral 0) (cbox a b)" assumes ne: "box a b \ {}" assumes x: "x \ cbox a b" shows "f x = 0" proof - have "0 < emeasure lborel (box a b)" using ne x unfolding emeasure_lborel_box_eq by (force intro!: prod_pos simp: mem_box algebra_simps) then show ?thesis using assms by (intro has_integral_0_closure_imp_0[of "box a b" f x]) (auto simp: emeasure_lborel_box_eq emeasure_lborel_cbox_eq algebra_simps mem_box) qed subsection\Various common equivalent forms of function measurability\ lemma indicator_sum_eq: fixes m::real and f :: "'a \ real" assumes "\m\ \ 2 ^ (2*n)" "m/2^n \ f x" "f x < (m+1)/2^n" "m \ \" shows "(\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) = m/2^n" (is "sum ?f ?S = _") proof - have "sum ?f ?S = sum (\k. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {m}" proof (rule comm_monoid_add_class.sum.mono_neutral_right) show "finite ?S" by (rule finite_abs_int_segment) show "{m} \ {k \ \. \k\ \ 2 ^ (2*n)}" using assms by auto show "\i\{k \ \. \k\ \ 2 ^ (2*n)} - {m}. ?f i = 0" using assms by (auto simp: indicator_def Ints_def abs_le_iff field_split_simps) qed also have "\ = m/2^n" using assms by (auto simp: indicator_def not_less) finally show ?thesis . qed lemma measurable_on_sf_limit_lemma1: fixes f :: "'a::euclidean_space \ real" assumes "\a b. {x \ S. a \ f x \ f x < b} \ sets (lebesgue_on S)" obtains g where "\n. g n \ borel_measurable (lebesgue_on S)" "\n. finite(range (g n))" "\x. (\n. g n x) \ f x" proof show "(\x. sum (\k::real. k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1)/2^n} x) {k \ \. \k\ \ 2 ^ (2*n)}) \ borel_measurable (lebesgue_on S)" (is "?g \ _") for n proof - have "\k. \k \ \; \k\ \ 2 ^ (2*n)\ \ Measurable.pred (lebesgue_on S) (\x. k / (2^n) \ f x \ f x < (k+1) / (2^n))" using assms by (force simp: pred_def space_restrict_space) then show ?thesis by (simp add: field_class.field_divide_inverse) qed show "finite (range (?g n))" for n proof - have "range (?g n) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof clarify fix x show "?g n x \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" proof (cases "\k::real. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ (f x) \ (f x) < (k+1)/2^n") case True then show ?thesis apply clarify by (subst indicator_sum_eq) auto next case False then have "?g n x = 0" by auto then show ?thesis by force qed qed moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (simp add: finite_abs_int_segment) ultimately show ?thesis using finite_subset by blast qed show "(\n. ?g n x) \ f x" for x proof (rule LIMSEQ_I) fix e::real assume "e > 0" obtain N1 where N1: "\f x\ < 2 ^ N1" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by force have "norm (?g n x - f x) < e" if n: "n \ max N1 N2" for n proof - define m where "m \ floor(2^n * (f x))" have "1 \ \2^n\ * e" using n N2 \e > 0\ less_eq_real_def less_le_trans by (fastforce simp add: field_split_simps) then have *: "\x \ y; y < x + 1\ \ abs(x - y) < \2^n\ * e" for x y::real by linarith have "\2^n\ * \m/2^n - f x\ = \2^n * (m/2^n - f x)\" by (simp add: abs_mult) also have "\ = \real_of_int \2^n * f x\ - f x * 2^n\" by (simp add: algebra_simps m_def) also have "\ < \2^n\ * e" by (rule *; simp add: mult.commute) finally have "\2^n\ * \m/2^n - f x\ < \2^n\ * e" . then have me: "\m/2^n - f x\ < e" by simp have "\real_of_int m\ \ 2 ^ (2*n)" proof (cases "f x < 0") case True then have "-\f x\ \ \(2::real) ^ N1\" using N1 le_floor_iff minus_le_iff by fastforce with n True have "\real_of_int\f x\\ \ 2 ^ N1" by linarith also have "\ \ 2^n" using n by (simp add: m_def) finally have "\real_of_int \f x\\ * 2^n \ 2^n * 2^n" by simp moreover have "\real_of_int \2^n * f x\\ \ \real_of_int \f x\\ * 2^n" proof - have "\real_of_int \2^n * f x\\ = - (real_of_int \2^n * f x\)" using True by (simp add: abs_if mult_less_0_iff) also have "\ \ - (real_of_int (\(2::real) ^ n\ * \f x\))" using le_mult_floor_Ints [of "(2::real)^n"] by simp also have "\ \ \real_of_int \f x\\ * 2^n" using True by simp finally show ?thesis . qed ultimately show ?thesis by (metis (no_types, hide_lams) m_def order_trans power2_eq_square power_even_eq) next case False with n N1 have "f x \ 2^n" by (simp add: not_less) (meson less_eq_real_def one_le_numeral order_trans power_increasing) moreover have "0 \ m" using False m_def by force ultimately show ?thesis by (metis abs_of_nonneg floor_mono le_floor_iff m_def of_int_0_le_iff power2_eq_square power_mult real_mult_le_cancel_iff1 zero_less_numeral mult.commute zero_less_power) qed then have "?g n x = m/2^n" by (rule indicator_sum_eq) (auto simp add: m_def field_split_simps, linarith) then have "norm (?g n x - f x) = norm (m/2^n - f x)" by simp also have "\ < e" by (simp add: me) finally show ?thesis . qed then show "\no. \n\no. norm (?g n x - f x) < e" by blast qed qed lemma borel_measurable_simple_function_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x))" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite (range (g n))) \ (\x. (\n. g n x) \ f x)" if f: "\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S)" proof - have "\g. (\n. (g n) \ borel_measurable (lebesgue_on S)) \ (\n. finite(image (g n) UNIV)) \ (\x. ((\n. g n x) \ f x \ i))" if "i \ Basis" for i proof (rule measurable_on_sf_limit_lemma1 [of S "\x. f x \ i"]) show "{x \ S. a \ f x \ i \ f x \ i < b} \ sets (lebesgue_on S)" for a b proof - have "{x \ S. a \ f x \ i \ f x \ i < b} = {x \ S. f x \ i < b} - {x \ S. a > f x \ i}" by auto also have "\ \ sets (lebesgue_on S)" using f that by blast finally show ?thesis . qed qed blast then obtain g where g: "\i n. i \ Basis \ g i n \ borel_measurable (lebesgue_on S)" "\i n. i \ Basis \ finite(range (g i n))" "\i x. i \ Basis \ ((\n. g i n x) \ f x \ i)" by metis show ?thesis proof (intro conjI allI exI) show "(\x. \i\Basis. g i n x *\<^sub>R i) \ borel_measurable (lebesgue_on S)" for n by (intro borel_measurable_sum borel_measurable_scaleR) (auto intro: g) show "finite (range (\x. \i\Basis. g i n x *\<^sub>R i))" for n proof - have "range (\x. \i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` PiE Basis (\i. range (g i n))" proof clarify fix x show "(\i\Basis. g i n x *\<^sub>R i) \ (\h. \i\Basis. h i *\<^sub>R i) ` (\\<^sub>E i\Basis. range (g i n))" by (rule_tac x="\i\Basis. g i n x" in image_eqI) auto qed moreover have "finite(PiE Basis (\i. range (g i n)))" by (simp add: g finite_PiE) ultimately show ?thesis by (metis (mono_tags, lifting) finite_surj) qed show "(\n. \i\Basis. g i n x *\<^sub>R i) \ f x" for x proof - have "(\n. \i\Basis. g i n x *\<^sub>R i) \ (\i\Basis. (f x \ i) *\<^sub>R i)" by (auto intro!: tendsto_sum tendsto_scaleR g) moreover have "(\i\Basis. (f x \ i) *\<^sub>R i) = f x" using euclidean_representation by blast ultimately show ?thesis by metis qed qed qed moreover have "f \ borel_measurable (lebesgue_on S)" if meas_g: "\n. g n \ borel_measurable (lebesgue_on S)" and fin: "\n. finite (range (g n))" and to_f: "\x. (\n. g n x) \ f x" for g by (rule borel_measurable_LIMSEQ_metric [OF meas_g to_f]) ultimately show ?thesis using borel_measurable_vimage_halfspace_component_lt by blast qed subsection \Lebesgue sets and continuous images\ proposition lebesgue_regular_inner: assumes "S \ sets lebesgue" obtains K C where "negligible K" "\n::nat. compact(C n)" "S = (\n. C n) \ K" proof - have "\T. closed T \ T \ S \ (S - T) \ lmeasurable \ emeasure lebesgue (S - T) < ennreal ((1/2)^n)" for n using sets_lebesgue_inner_closed assms by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) then obtain C where clo: "\n. closed (C n)" and subS: "\n. C n \ S" and mea: "\n. (S - C n) \ lmeasurable" and less: "\n. emeasure lebesgue (S - C n) < ennreal ((1/2)^n)" by metis have "\F. (\n::nat. compact(F n)) \ (\n. F n) = C m" for m::nat by (metis clo closed_Union_compact_subsets) then obtain D :: "[nat,nat] \ 'a set" where D: "\m n. compact(D m n)" "\m. (\n. D m n) = C m" by metis let ?C = "from_nat_into (\m. range (D m))" have "countable (\m. range (D m))" by blast have "range (from_nat_into (\m. range (D m))) = (\m. range (D m))" using range_from_nat_into by simp then have CD: "\m n. ?C k = D m n" for k by (metis (mono_tags, lifting) UN_iff rangeE range_eqI) show thesis proof show "negligible (S - (\n. C n))" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" then obtain n where n: "(1/2)^n < e" using real_arch_pow_inv [of e "1/2"] by auto show "\T. S - (\n. C n) \ T \ T \ lmeasurable \ measure lebesgue T \ e" proof (intro exI conjI) show "S - (\n. C n) \ S - C n" by blast show "S - C n \ lmeasurable" by (simp add: mea) show "measure lebesgue (S - C n) \ e" using less [of n] n by (simp add: emeasure_eq_measure2 less_le mea) qed qed show "compact (?C n)" for n using CD D by metis show "S = (\n. ?C n) \ (S - (\n. C n))" (is "_ = ?rhs") proof show "S \ ?rhs" using D by fastforce show "?rhs \ S" using subS D CD by auto (metis Sup_upper range_eqI subsetCE) qed qed qed lemma sets_lebesgue_continuous_image: assumes T: "T \ sets lebesgue" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" and "T \ S" shows "f ` T \ sets lebesgue" proof - obtain K C where "negligible K" and com: "\n::nat. compact(C n)" and Teq: "T = (\n. C n) \ K" using lebesgue_regular_inner [OF T] by metis then have comf: "\n::nat. compact(f ` C n)" by (metis Un_subset_iff Union_upper \T \ S\ compact_continuous_image contf continuous_on_subset rangeI) have "((\n. f ` C n) \ f ` K) \ sets lebesgue" proof (rule sets.Un) have "K \ S" using Teq \T \ S\ by blast show "(\n. f ` C n) \ sets lebesgue" proof (rule sets.countable_Union) show "range (\n. f ` C n) \ sets lebesgue" using borel_compact comf by (auto simp: borel_compact) qed auto show "f ` K \ sets lebesgue" by (simp add: \K \ S\ \negligible K\ negim negligible_imp_sets) qed then show ?thesis by (simp add: Teq image_Un image_Union) qed lemma differentiable_image_in_sets_lebesgue: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f`S \ sets lebesgue" proof (rule sets_lebesgue_continuous_image [OF S]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed auto lemma sets_lebesgue_on_continuous_image: assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" shows "f ` X \ sets (lebesgue_on (f ` S))" proof - have "X \ S" by (metis S X sets.Int_space_eq2 sets_restrict_space_iff) moreover have "f ` S \ sets lebesgue" using S contf negim sets_lebesgue_continuous_image by blast moreover have "f ` X \ sets lebesgue" by (metis S X contf negim sets_lebesgue_continuous_image sets_restrict_space_iff space_restrict_space space_restrict_space2) ultimately show ?thesis by (auto simp: sets_restrict_space_iff) qed lemma differentiable_image_in_sets_lebesgue_on: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f ` X \ sets (lebesgue_on (f`S))" proof (rule sets_lebesgue_on_continuous_image [OF S X]) show "continuous_on S f" by (meson differentiable_imp_continuous_on f) show "\T. \negligible T; T \ S\ \ negligible (f ` T)" using differentiable_on_subset f by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed subsection \Affine lemmas\ lemma borel_measurable_affine: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes f: "f \ borel_measurable lebesgue" and "c \ 0" shows "(\x. f(t + c *\<^sub>R x)) \ borel_measurable lebesgue" proof - { fix a b have "{x. f x \ cbox a b} \ sets lebesgue" using f cbox_borel lebesgue_measurable_vimage_borel by blast then have "(\x. (x - t) /\<^sub>R c) ` {x. f x \ cbox a b} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "(\x. (x - t) /\<^sub>R c) differentiable_on {x. f x \ cbox a b}" unfolding differentiable_on_def differentiable_def by (rule \c \ 0\ derivative_eq_intros strip exI | simp)+ qed auto moreover have "{x. f(t + c *\<^sub>R x) \ cbox a b} = (\x. (x-t) /\<^sub>R c) ` {x. f x \ cbox a b}" using \c \ 0\ by (auto simp: image_def) ultimately have "{x. f(t + c *\<^sub>R x) \ cbox a b} \ sets lebesgue" by (auto simp: borel_measurable_vimage_closed_interval) } then show ?thesis by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) qed lemma lebesgue_integrable_real_affine: fixes f :: "real \ 'a :: euclidean_space" assumes f: "integrable lebesgue f" and "c \ 0" shows "integrable lebesgue (\x. f(t + c * x))" proof - have "(\x. norm (f x)) \ borel_measurable lebesgue" by (simp add: borel_measurable_integrable f) then show ?thesis using assms borel_measurable_affine [of f c] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) qed lemma lebesgue_integrable_real_affine_iff: fixes f :: "real \ 'a :: euclidean_space" shows "c \ 0 \ integrable lebesgue (\x. f(t + c * x)) \ integrable lebesgue f" using lebesgue_integrable_real_affine[of f c t] lebesgue_integrable_real_affine[of "\x. f(t + c * x)" "1/c" "-t/c"] by (auto simp: field_simps) lemma\<^marker>\tag important\ lebesgue_integral_real_affine: fixes f :: "real \ 'a :: euclidean_space" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lebesgue) = \c\ *\<^sub>R (\x. f(t + c * x) \lebesgue)" proof cases have "(\x. t + c * x) \ lebesgue \\<^sub>M lebesgue" using lebesgue_affine_measurable[where c= "\x::real. c"] \c \ 0\ by simp moreover assume "integrable lebesgue f" ultimately show ?thesis by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) next assume "\ integrable lebesgue f" with c show ?thesis by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) qed lemma has_bochner_integral_lebesgue_real_affine_iff: fixes i :: "'a :: euclidean_space" shows "c \ 0 \ has_bochner_integral lebesgue f i \ has_bochner_integral lebesgue (\x. f(t + c * x)) (i /\<^sub>R \c\)" unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma has_bochner_integral_reflect_real_lemma[intro]: fixes f :: "real \ 'a::euclidean_space" assumes "has_bochner_integral (lebesgue_on {a..b}) f i" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i" proof - have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x by (auto simp: indicator_def) have i: "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" using assms by (auto simp: has_bochner_integral_restrict_space) then have "has_bochner_integral lebesgue (\x. indicator {-b..-a} x *\<^sub>R f(-x)) i" using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\x. indicator {a..b} x *\<^sub>R f x)" i 0] by (auto simp: eq) then show ?thesis by (auto simp: has_bochner_integral_restrict_space) qed lemma has_bochner_integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i \ has_bochner_integral (lebesgue_on {a..b}) f i" by (auto simp: dest: has_bochner_integral_reflect_real_lemma) lemma integrable_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integrable (lebesgue_on {-b..-a}) (\x. f(-x)) \ integrable (lebesgue_on {a..b}) f" by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) lemma integral_reflect_real[simp]: fixes f :: "real \ 'a::euclidean_space" shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f" using has_bochner_integral_reflect_real [of b a f] by (metis has_bochner_integral_iff not_integrable_integral_eq) subsection\More results on integrability\ lemma integrable_on_all_intervals_UNIV: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. f integrable_on cbox a b" and normf: "\x. norm(f x) \ g x" and g: "g integrable_on UNIV" shows "f integrable_on UNIV" proof - have intg: "(\a b. g integrable_on cbox a b)" and gle_e: "\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ \integral (cbox a b) g - integral (cbox c d) g\ < e" using g by (auto simp: integrable_alt_subset [of _ UNIV] intf) have le: "norm (integral (cbox a b) f - integral (cbox c d) f) \ \integral (cbox a b) g - integral (cbox c d) g\" if "cbox a b \ cbox c d" for a b c d proof - have "norm (integral (cbox a b) f - integral (cbox c d) f) = norm (integral (cbox c d - cbox a b) f)" using intf that by (simp add: norm_minus_commute integral_setdiff) also have "\ \ integral (cbox c d - cbox a b) g" proof (rule integral_norm_bound_integral [OF _ _ normf]) show "f integrable_on cbox c d - cbox a b" "g integrable_on cbox c d - cbox a b" by (meson integrable_integral integrable_setdiff intf intg negligible_setdiff that)+ qed also have "\ = integral (cbox c d) g - integral (cbox a b) g" using intg that by (simp add: integral_setdiff) also have "\ \ \integral (cbox a b) g - integral (cbox c d) g\" by simp finally show ?thesis . qed show ?thesis using gle_e apply (simp add: integrable_alt_subset [of _ UNIV] intf) apply (erule imp_forward all_forward ex_forward asm_rl)+ by (meson not_less order_trans le) qed lemma integrable_on_all_intervals_integrable_bound: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b" and normf: "\x. x \ S \ norm(f x) \ g x" and g: "g integrable_on S" shows "f integrable_on S" using integrable_on_all_intervals_UNIV [OF intf, of "(\x. if x \ S then g x else 0)"] by (simp add: g integrable_restrict_UNIV normf) lemma measurable_bounded_lemma: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable lebesgue" and g: "g integrable_on cbox a b" and normf: "\x. x \ cbox a b \ norm(f x) \ g x" shows "f integrable_on cbox a b" proof - have "g absolutely_integrable_on cbox a b" by (metis (full_types) add_increasing g le_add_same_cancel1 nonnegative_absolutely_integrable_1 norm_ge_zero normf) then have "integrable (lebesgue_on (cbox a b)) g" by (simp add: integrable_restrict_space set_integrable_def) then have "integrable (lebesgue_on (cbox a b)) f" proof (rule Bochner_Integration.integrable_bound) show "AE x in lebesgue_on (cbox a b). norm (f x) \ norm (g x)" by (rule AE_I2) (auto intro: normf order_trans) qed (simp add: f measurable_restrict_space1) then show ?thesis by (simp add: integrable_on_lebesgue_on) qed proposition measurable_bounded_by_integrable_imp_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "g integrable_on S" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "f integrable_on S" proof (rule integrable_on_all_intervals_integrable_bound [OF _ normf g]) show "(\x. if x \ S then f x else 0) integrable_on cbox a b" for a b proof (rule measurable_bounded_lemma) show "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" by (simp add: S borel_measurable_if f) show "(\x. if x \ S then g x else 0) integrable_on cbox a b" by (simp add: g integrable_altD(1)) show "norm (if x \ S then f x else 0) \ (if x \ S then g x else 0)" for x using normf by simp qed qed lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" shows "integrable (lebesgue_on S) f" proof - have "f absolutely_integrable_on S" by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) then show ?thesis by (simp add: S integrable_restrict_space set_integrable_def) qed lemma measurable_bounded_by_integrable_imp_integrable_real: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable (lebesgue_on S)" "g integrable_on S" "\x. x \ S \ abs(f x) \ g x" "S \ sets lebesgue" shows "f integrable_on S" using measurable_bounded_by_integrable_imp_integrable [of f S g] assms by simp subsection\ Relation between Borel measurability and integrability.\ lemma integrable_imp_measurable_weak: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" by (metis (mono_tags, lifting) assms has_integral_implies_lebesgue_measurable borel_measurable_restrict_space_iff integrable_on_def sets.Int_space_eq2) lemma integrable_imp_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on S" shows "f \ borel_measurable (lebesgue_on S)" proof - have "(UNIV::'a set) \ sets lborel" by simp then show ?thesis by (metis (mono_tags, lifting) assms borel_measurable_if_D integrable_imp_measurable_weak integrable_restrict_UNIV lebesgue_on_UNIV_eq sets_lebesgue_on_refl) qed lemma integrable_iff_integrable_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "(\\<^sup>+ x. ennreal (norm (f x)) \lebesgue_on S) < \" shows "integrable (lebesgue_on S) f \ f integrable_on S" using assms integrable_iff_bounded integrable_imp_measurable integrable_on_lebesgue_on by blast lemma absolutely_integrable_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (norm \ f)" (is "?lhs = ?rhs") proof assume L: ?lhs then have "f \ borel_measurable (lebesgue_on S)" by (simp add: absolutely_integrable_on_def integrable_imp_measurable) then show ?rhs using assms set_integrable_norm [of lebesgue S f] L by (simp add: integrable_restrict_space set_integrable_def) next assume ?rhs then show ?lhs using assms integrable_on_lebesgue_on by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) qed lemma absolutely_integrable_measurable_real: fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. \f x\)" by (simp add: absolutely_integrable_measurable assms o_def) lemma absolutely_integrable_measurable_real': fixes f :: "'a::euclidean_space \ real" assumes "S \ sets lebesgue" shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ (\x. \f x\) integrable_on S" using assms apply (auto simp: absolutely_integrable_measurable integrable_on_lebesgue_on) apply (simp add: integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_lebesgue_integrable) using abs_absolutely_integrableI_1 absolutely_integrable_measurable measurable_bounded_by_integrable_imp_integrable_real by blast lemma absolutely_integrable_imp_borel_measurable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f absolutely_integrable_on S" "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using absolutely_integrable_measurable assms by blast lemma measurable_bounded_by_integrable_imp_absolutely_integrable: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" and "g integrable_on S" and "\x. x \ S \ norm(f x) \ (g x)" shows "f absolutely_integrable_on S" using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast proposition negligible_differentiable_vimage: fixes f :: "'a \ 'a::euclidean_space" assumes "negligible T" and f': "\x. x \ S \ inj(f' x)" and derf: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "negligible {x \ S. f x \ T}" proof - define U where "U \ \n::nat. {x \ S. \y. y \ S \ norm(y - x) < 1/n \ norm(y - x) \ n * norm(f y - f x)}" have "negligible {x \ U n. f x \ T}" if "n > 0" for n proof (subst locally_negligible_alt, clarify) fix a assume a: "a \ U n" and fa: "f a \ T" define V where "V \ {x. x \ U n \ f x \ T} \ ball a (1 / n / 2)" show "\V. openin (top_of_set {x \ U n. f x \ T}) V \ a \ V \ negligible V" proof (intro exI conjI) have noxy: "norm(x - y) \ n * norm(f x - f y)" if "x \ V" "y \ V" for x y using that unfolding U_def V_def mem_Collect_eq Int_iff mem_ball dist_norm by (meson norm_triangle_half_r) then have "inj_on f V" by (force simp: inj_on_def) then obtain g where g: "\x. x \ V \ g(f x) = x" by (metis inv_into_f_f) have "\T' B. open T' \ f x \ T' \ (\y\f ` V \ T \ T'. norm (g y - g (f x)) \ B * norm (y - f x))" if "f x \ T" "x \ V" for x apply (rule_tac x = "ball (f x) 1" in exI) using that noxy by (force simp: g) then have "negligible (g ` (f ` V \ T))" by (force simp: \negligible T\ negligible_Int intro!: negligible_locally_Lipschitz_image) moreover have "V \ g ` (f ` V \ T)" by (force simp: g image_iff V_def) ultimately show "negligible V" by (rule negligible_subset) qed (use a fa V_def that in auto) qed with negligible_countable_Union have "negligible (\n \ {0<..}. {x. x \ U n \ f x \ T})" by auto moreover have "{x \ S. f x \ T} \ (\n \ {0<..}. {x. x \ U n \ f x \ T})" proof clarsimp fix x assume "x \ S" and "f x \ T" then obtain inj: "inj(f' x)" and der: "(f has_derivative f' x) (at x within S)" using assms by metis moreover have "linear(f' x)" and eps: "\\. \ > 0 \ \\>0. \y\S. norm (y - x) < \ \ norm (f y - f x - f' x (y - x)) \ \ * norm (y - x)" using der by (auto simp: has_derivative_within_alt linear_linear) ultimately obtain g where "linear g" and g: "g \ f' x = id" using linear_injective_left_inverse by metis then obtain B where "B > 0" and B: "\z. B * norm z \ norm(f' x z)" using linear_invertible_bounded_below_pos \linear (f' x)\ by blast then obtain i where "i \ 0" and i: "1 / real i < B" by (metis (full_types) inverse_eq_divide real_arch_invD) then obtain \ where "\ > 0" and \: "\y. \y\S; norm (y - x) < \\ \ norm (f y - f x - f' x (y - x)) \ (B - 1 / real i) * norm (y - x)" using eps [of "B - 1/i"] by auto then obtain j where "j \ 0" and j: "inverse (real j) < \" using real_arch_inverse by blast have "norm (y - x)/(max i j) \ norm (f y - f x)" if "y \ S" and less: "norm (y - x) < 1 / (max i j)" for y proof - have "1 / real (max i j) < \" using j \j \ 0\ \0 < \\ by (auto simp: field_split_simps max_mult_distrib_left of_nat_max) then have "norm (y - x) < \" using less by linarith with \ \y \ S\ have le: "norm (f y - f x - f' x (y - x)) \ B * norm (y - x) - norm (y - x)/i" by (auto simp: algebra_simps) have *: "\norm(f - f' - y) \ b - c; b \ norm y; d \ c\ \ d \ norm(f - f')" for b c d and y f f'::'a using norm_triangle_ineq3 [of "f - f'" y] by simp show ?thesis apply (rule * [OF le B]) using \i \ 0\ \j \ 0\ by (simp add: field_split_simps max_mult_distrib_left of_nat_max less_max_iff_disj) qed with \x \ S\ \i \ 0\ \j \ 0\ show "\n\{0<..}. x \ U n" by (rule_tac x="max i j" in bexI) (auto simp: field_simps U_def less_max_iff_disj) qed ultimately show ?thesis by (rule negligible_subset) qed lemma absolutely_integrable_Un: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "f absolutely_integrable_on S" and T: "f absolutely_integrable_on T" shows "f absolutely_integrable_on (S \ T)" proof - have [simp]: "{x. (if x \ A then f x else 0) \ 0} = {x \ A. f x \ 0}" for A by auto let ?ST = "{x \ S. f x \ 0} \ {x \ T. f x \ 0}" have "?ST \ sets lebesgue" proof (rule Sigma_Algebra.sets.Int) have "f integrable_on S" using S absolutely_integrable_on_def by blast then have "(\x. if x \ S then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ S. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto next have "f integrable_on T" using T absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then show "{x \ T. f x \ 0} \ sets lebesgue" unfolding borel_measurable_vimage_open by (rule allE [where x = "-{0}"]) auto qed then have "f absolutely_integrable_on ?ST" by (rule set_integrable_subset [OF S]) auto then have Int: "(\x. if x \ ?ST then f x else 0) absolutely_integrable_on UNIV" using absolutely_integrable_restrict_UNIV by blast have "(\x. if x \ S then f x else 0) absolutely_integrable_on UNIV" "(\x. if x \ T then f x else 0) absolutely_integrable_on UNIV" using S T absolutely_integrable_restrict_UNIV by blast+ then have "(\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)) absolutely_integrable_on UNIV" by (rule set_integral_add) then have "(\x. ((if x \ S then f x else 0) + (if x \ T then f x else 0)) - (if x \ ?ST then f x else 0)) absolutely_integrable_on UNIV" using Int by (rule set_integral_diff) then have "(\x. if x \ S \ T then f x else 0) absolutely_integrable_on UNIV" by (rule absolutely_integrable_spike) (auto intro: empty_imp_negligible) then show ?thesis unfolding absolutely_integrable_restrict_UNIV . qed lemma absolutely_integrable_on_combine: fixes f :: "real \ 'a::euclidean_space" assumes "f absolutely_integrable_on {a..c}" and "f absolutely_integrable_on {c..b}" and "a \ c" and "c \ b" shows "f absolutely_integrable_on {a..b}" by (metis absolutely_integrable_Un assms ivl_disj_un_two_touch(4)) lemma uniform_limit_set_lebesgue_integral_at_top: fixes f :: "'a \ real \ 'b::{banach, second_countable_topology}" and g :: "real \ real" assumes bound: "\x y. x \ A \ y \ a \ norm (f x y) \ g y" assumes integrable: "set_integrable M {a..} g" assumes measurable: "\x. x \ A \ set_borel_measurable M {a..} (f x)" assumes "sets borel \ sets M" shows "uniform_limit A (\b x. LINT y:{a..b}|M. f x y) (\x. LINT y:{a..}|M. f x y) at_top" proof (cases "A = {}") case False then obtain x where x: "x \ A" by auto have g_nonneg: "g y \ 0" if "y \ a" for y proof - have "0 \ norm (f x y)" by simp also have "\ \ g y" using bound[OF x that] by simp finally show ?thesis . qed have integrable': "set_integrable M {a..} (\y. f x y)" if "x \ A" for x unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) show "integrable M (\x. indicator {a..} x * g x)" using integrable by (simp add: set_integrable_def) show "(\y. indicat_real {a..} y *\<^sub>R f x y) \ borel_measurable M" using measurable[OF that] by (simp add: set_borel_measurable_def) show "AE y in M. norm (indicat_real {a..} y *\<^sub>R f x y) \ norm (indicat_real {a..} y * g y)" using bound[OF that] by (intro AE_I2) (auto simp: indicator_def g_nonneg) qed show ?thesis proof (rule uniform_limitI) fix e :: real assume e: "e > 0" have sets [intro]: "A \ sets M" if "A \ sets borel" for A using that assms by blast have "((\b. LINT y:{a..b}|M. g y) \ (LINT y:{a..}|M. g y)) at_top" by (intro tendsto_set_lebesgue_integral_at_top assms sets) auto with e obtain b0 :: real where b0: "\b\b0. \(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: tendsto_iff eventually_at_top_linorder dist_real_def abs_minus_commute) define b where "b = max a b0" have "a \ b" by (simp add: b_def) from b0 have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" by (auto simp: b_def) also have "{a..} = {a..b} \ {b<..}" by (auto simp: b_def) also have "\(LINT y:\|M. g y) - (LINT y:{a..b}|M. g y)\ = \(LINT y:{b<..}|M. g y)\" using \a \ b\ by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "\(LINT y:{b<..}|M. g y)\ = (LINT y:{b<..}|M. g y)" by simp finally have less: "(LINT y:{b<..}|M. g y) < e" . have "eventually (\b. b \ b0) at_top" by (rule eventually_ge_at_top) moreover have "eventually (\b. b \ a) at_top" by (rule eventually_ge_at_top) ultimately show "eventually (\b. \x\A. dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e) at_top" proof eventually_elim case (elim b) show ?case proof fix x assume x: "x \ A" have "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) = norm ((LINT y:{a..}|M. f x y) - (LINT y:{a..b}|M. f x y))" by (simp add: dist_norm norm_minus_commute) also have "{a..} = {a..b} \ {b<..}" using elim by auto also have "(LINT y:\|M. f x y) - (LINT y:{a..b}|M. f x y) = (LINT y:{b<..}|M. f x y)" using elim x by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable']) also have "norm \ \ (LINT y:{b<..}|M. norm (f x y))" using elim x by (intro set_integral_norm_bound set_integrable_subset[OF integrable']) auto also have "\ \ (LINT y:{b<..}|M. g y)" using elim x bound g_nonneg by (intro set_integral_mono set_integrable_norm set_integrable_subset[OF integrable'] set_integrable_subset[OF integrable]) auto also have "(LINT y:{b<..}|M. g y) \ 0" using g_nonneg \a \ b\ unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg) (auto simp: indicator_def) hence "(LINT y:{b<..}|M. g y) = \(LINT y:{b<..}|M. g y)\" by simp also have "\ = \(LINT y:{a..b} \ {b<..}|M. g y) - (LINT y:{a..b}|M. g y)\" using elim by (subst set_integral_Un) (auto intro!: set_integrable_subset[OF integrable]) also have "{a..b} \ {b<..} = {a..}" using elim by auto also have "\(LINT y:{a..}|M. g y) - (LINT y:{a..b}|M. g y)\ < e" using b0 elim by blast finally show "dist (LINT y:{a..b}|M. f x y) (LINT y:{a..}|M. f x y) < e" . qed qed qed qed auto subsubsection\Differentiability of inverse function (most basic form)\ proposition has_derivative_inverse_within: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes der_f: "(f has_derivative f') (at a within S)" and cont_g: "continuous (at (f a) within f ` S) g" and "a \ S" "linear g'" and id: "g' \ f' = id" and gf: "\x. x \ S \ g(f x) = x" shows "(g has_derivative g') (at (f a) within f ` S)" proof - have [simp]: "g' (f' x) = x" for x by (simp add: local.id pointfree_idE) have "bounded_linear f'" and f': "\e. e>0 \ \d>0. \y\S. norm (y - a) < d \ norm (f y - f a - f' (y - a)) \ e * norm (y - a)" using der_f by (auto simp: has_derivative_within_alt) obtain C where "C > 0" and C: "\x. norm (g' x) \ C * norm x" using linear_bounded_pos [OF \linear g'\] by metis obtain B k where "B > 0" "k > 0" and Bk: "\x. \x \ S; norm(f x - f a) < k\ \ norm(x - a) \ B * norm(f x - f a)" proof - obtain B where "B > 0" and B: "\x. B * norm x \ norm (f' x)" using linear_inj_bounded_below_pos [of f'] \linear g'\ id der_f has_derivative_linear linear_invertible_bounded_below_pos by blast then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ B / 2 * norm (y - a)" using f' [of "B/2"] by auto then obtain e where "e > 0" and e: "\x. \x \ S; norm (f x - f a) < e\ \ norm (g (f x) - g (f a)) < d" using cont_g by (auto simp: continuous_within_eps_delta dist_norm) show thesis proof show "2/B > 0" using \B > 0\ by simp show "norm (x - a) \ 2 / B * norm (f x - f a)" if "x \ S" "norm (f x - f a) < e" for x proof - have xa: "norm (x - a) < d" using e [OF that] gf by (simp add: \a \ S\ that) have *: "\norm(y - f') \ B / 2 * norm x; B * norm x \ norm f'\ \ norm y \ B / 2 * norm x" for y f'::'b and x::'a using norm_triangle_ineq3 [of y f'] by linarith show ?thesis using * [OF d [OF \x \ S\ xa] B] \B > 0\ by (simp add: field_simps) qed qed (use \e > 0\ in auto) qed show ?thesis unfolding has_derivative_within_alt proof (intro conjI impI allI) show "bounded_linear g'" using \linear g'\ by (simp add: linear_linear) next fix e :: "real" assume "e > 0" then obtain d where "d>0" and d: "\y. \y \ S; norm (y - a) < d\ \ norm (f y - f a - f' (y - a)) \ e / (B * C) * norm (y - a)" using f' [of "e / (B * C)"] \B > 0\ \C > 0\ by auto have "norm (x - a - g' (f x - f a)) \ e * norm (f x - f a)" if "x \ S" and lt_k: "norm (f x - f a) < k" and lt_dB: "norm (f x - f a) < d/B" for x proof - have "norm (x - a) \ B * norm(f x - f a)" using Bk lt_k \x \ S\ by blast also have "\ < d" by (metis \0 < B\ lt_dB mult.commute pos_less_divide_eq) finally have lt_d: "norm (x - a) < d" . have "norm (x - a - g' (f x - f a)) \ norm(g'(f x - f a - (f' (x - a))))" by (simp add: linear_diff [OF \linear g'\] norm_minus_commute) also have "\ \ C * norm (f x - f a - f' (x - a))" using C by blast also have "\ \ e * norm (f x - f a)" proof - have "norm (f x - f a - f' (x - a)) \ e / (B * C) * norm (x - a)" using d [OF \x \ S\ lt_d] . also have "\ \ (norm (f x - f a) * e) / C" using \B > 0\ \C > 0\ \e > 0\ by (simp add: field_simps Bk lt_k \x \ S\) finally show ?thesis using \C > 0\ by (simp add: field_simps) qed finally show ?thesis . qed then show "\d>0. \y\f ` S. norm (y - f a) < d \ norm (g y - g (f a) - g' (y - f a)) \ e * norm (y - f a)" apply (rule_tac x="min k (d / B)" in exI) using \k > 0\ \B > 0\ \d > 0\ \a \ S\ by (auto simp: gf) qed qed end diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4175 +1,4175 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis unfolding path_def apply safe apply (subst g) apply (rule continuous_on_compose) apply (auto intro: continuous_intros) done qed lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using *[of "reversepath g"] *[of g] unfolding reversepath_reversepath by (rule iffI) qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: "pathfinish g1 = pathstart g2 \ path_image(g1 +++ g2) = path_image g1 \ path_image g2" apply (rule subset_antisym [OF path_image_join_subset]) apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def) apply (drule sym) apply (rule_tac x="xa/2" in bexI, auto) apply (rule ccontr) apply (drule_tac x="(xa+1)/2" in bspec) apply (auto simp: field_simps) apply (drule_tac x="1/2" in bspec, auto) done lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def) apply (metis eq_iff le_less_linear) apply (metis leD linear) using less_eq_real_def zero_le_one apply blast using less_eq_real_def zero_le_one apply blast done lemma connected_simple_path_endless: "simple_path c \ connected(path_image c - {pathstart c,pathfinish c})" apply (simp add: simple_path_endless) apply (rule connected_continuous_image) apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path) by auto lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by (auto simp: path_image_def reversepath_def) lemma path_imp_reversepath: "path g \ path(reversepath g)" apply (auto simp: path_def reversepath_def) using continuous_on_compose [of "{0..1}" "\x. 1 - x" g] apply (auto simp: continuous_on_op_minus) done lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have *: "{0..1::real} = {0..1/2} \ {1/2..1}" by auto have gg: "g2 0 = g1 1" by (metis assms(3) pathfinish_def pathstart_def) have 1: "continuous_on {0..1/2} (g1 +++ g2)" apply (rule continuous_on_eq [of _ "g1 \ (\x. 2*x)"]) apply (rule continuous_intros | simp add: joinpaths_def assms)+ done have "continuous_on {1/2..1} (g2 \ (\x. 2*x-1))" apply (rule continuous_on_subset [of "{1/2..1}"]) apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+ done then have 2: "continuous_on {1/2..1} (g1 +++ g2)" apply (rule continuous_on_eq [of "{1/2..1}" "g2 \ (\x. 2*x-1)"]) apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+ done show ?thesis apply (subst *) apply (rule continuous_on_closed_Un) using 1 2 apply auto done qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by (simp) lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xyI: "x = 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy apply auto apply (drule inj_onD [OF injg1]) using g21 [symmetric] xyI apply (auto dest: inj_onD [OF injg2]) done } note * = this { fix x and y::real assume xy: "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" "g1 (2 * x) = g2 (2 * y - 1)" have g1im: "g1 (2 * x) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x" in image_eqI, auto) done have "x = 0 \ y = 1" using subsetD [OF sb g1im] xy apply auto apply (force dest: inj_onD [OF injg1]) using g21 [symmetric] apply (auto dest: inj_onD [OF injg2]) done } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def path_join, clarify) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis **) apply (force dest: inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy by (auto dest: inj_onD [OF injg2]) } note * = this show ?thesis apply (simp add: arc_def inj_on_def) apply (clarsimp simp add: arc_imp_path assms) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis *) apply (force dest: inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" using assms(2) unfolding path_def continuous_on_iff apply (drule_tac x=0 in bspec, simp) by (metis half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec) apply (force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ apply (rule simple_path_joinE [OF arc_imp_simple_path assms]) using n1 by force next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" apply (rule continuous_intros | simp)+ apply (simp add: image_affinity_atLeastAtMost [where c=u]) using assms apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" apply (rule arc_simple_path_subpath) apply (force simp: simple_path_def)+ done lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath) apply (auto simp: path_image_def) done lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) then have com: "compact ({0..1} \ {u. g u \ closure (- S)})" apply (simp add: Int_commute [of "{0..1}"] compact_eq_bounded_closed closed_vimage_Int [unfolded vimage_def]) using compact_eq_bounded_closed apply fastforce done have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce { assume "u \ 0" then have "u > 0" using \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then have "g u \ closure S" by (simp add: frontier_def closure_approachable) } then show ?thesis apply (rule_tac u=u in that) apply (auto simp: \0 \ u\ \u \ 1\ gu interior_closure umin) using \_ \ 1\ interior_closure umin apply fastforce done qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (simp add: gunot) using \0 \ u\ u0 by (force simp: subpath_def gxin) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" using subpath_to_frontier_strong [OF g g1] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) using \0 \ u\ g0 disj apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) apply (drule_tac x="y/u" in spec) apply (auto split: if_split_asm) done qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis apply (rule that [of "subpath 0 u g"]) using assms u apply (simp_all add: path_image_subpath) apply (simp add: pathstart_def) apply (force simp: closed_segment_eq_real_ivl path_image_def) done qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis apply (rule that [OF \path h\]) using assms h apply auto apply (metis Diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff) done qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" by (auto simp: shiftpath_def) lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a apply (force simp: field_simps)+ done next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto lemma linepath_in_path: shows "x \ {0..1} \ linepath a b x \ closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_in_convex_hull: fixes x::real assumes a: "a \ convex hull s" and b: "b \ convex hull s" and x: "0\x" "x\1" shows "linepath a b x \ convex hull s" apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) using x apply (auto simp: linepath_image_01 [symmetric]) done lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" by (simp add: linepath_def) lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" by (simp add: linepath_def) lemma bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within s)" apply (simp add: linepath_def has_vector_derivative_def algebra_simps) apply (rule derivative_eq_intros | simp)+ done subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set s x \ Collect (path_component s x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component s x y" shows "x \ s" and "y \ s" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ s" shows "path_component s x x" unfolding path_defs apply (rule_tac x="\u. x" in exI) using assms apply (auto intro!: continuous_intros) done lemma path_component_refl_eq: "path_component s x x \ x \ s" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component s x y \ path_component s y x" unfolding path_component_def apply (erule exE) apply (rule_tac x="reversepath g" in exI, auto) done lemma path_component_trans: assumes "path_component s x y" and "path_component s y z" shows "path_component s x z" using assms unfolding path_component_def apply (elim exE) apply (rule_tac x="g +++ ga" in exI) apply (auto simp: path_image_join) done lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" unfolding path_component_def by auto lemma path_component_linepath: fixes s :: "'a::real_normed_vector set" shows "closed_segment a b \ s \ path_component s a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set s x = {y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set s x \ s" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set s x = {} \ x \ s" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "s \ t \ (path_component_set s x) \ (path_component_set t x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set s x \ path_component_set s y = path_component_set s x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected s \ (\x\s. \y\s. \g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected s \ (\x\s. path_component_set s x = s)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ t; path_connected t; t \ s\ \ t \ (path_component_set s x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes s :: "'a::real_normed_vector set" assumes "convex s" shows "path_connected s" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms[unfolded open_contains_ball] by auto have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' apply (rule_tac x="f \ g" in exI) unfolding path_defs apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)]) apply auto done qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def apply clarify apply (rule_tac x="\x. a" in exI) apply (simp add: image_constant_conv) apply (simp add: path_def) done lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on {0..1} (p \ ((*) y))" apply (rule continuous_intros)+ using p [unfolded path_def] y apply (auto simp: mult_le_one intro: continuous_on_subset [of _ p]) done then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def, clarify) apply (rule_tac x=g in exI) by (simp add: pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis apply (rule subset_antisym) apply (simp add: path_component_subset) by (simp add: True path_component_maximal path_component_refl) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis apply (rule connected_component_maximal) apply (auto simp: True path_component_subset path_component_refl path_connected_imp_connected) done next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" unfolding path_connected_space_def Ball_def apply (intro all_cong1 imp_cong refl, safe) using path_connectedin_path_image apply fastforce by (meson path_connectedin) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" using linear_injective_left_inverse [OF assms] apply clarify apply (rule_tac g=g in that) using assms apply (auto simp: homeomorphism_def eq_id_iff [symmetric] image_comp comp_def linear_conv_bounded_linear linear_continuous_on) done lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = h]) apply (rule continuous_intros | force)+ done moreover have "path (\z. (g z, y2))" using \path g\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = g]) apply (rule continuous_intros | force)+ done ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" apply (intro exI conjI) apply (rule 1) apply (rule 2) apply (metis hs pathstart_def pathstart_join) by (metis gf pathfinish_def pathfinish_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" apply (auto simp: disjnt_def) using path_component_eq apply fastforce using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" apply (rule iffI, metis (no_types) path_component_mem(1) path_component_refl) apply (erule disjE, metis Collect_empty_eq_bot path_component_eq_empty) apply (rule ext) apply (metis path_component_trans path_component_sym) done lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" apply (rule subset_antisym) using assms apply (metis mem_Collect_eq subsetCE path_component_eq_eq path_component_subset path_connected_path_component) by (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" apply (auto simp: path_component_in_topspace) apply (force simp: path_component_of_def pathin_const) done lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose) apply (force simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" apply (auto simp: path_component_of_def) using path_connectedin_path_image apply fastforce apply (metis path_connectedin) done lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "\y. path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) then show ?thesis apply (simp add: path_connectedin_def path_component_of_subset_topspace path_connected_space_iff_path_component) by (metis Int_absorb1 mem_Collect_eq path_component_of_equiv path_component_of_subset_topspace topspace_subtopology) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: "C \ path_components_of X \ (C \ {})" apply (clarsimp simp: path_components_of_def path_component_of_eq_empty) by (meson path_component_of_refl) lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" apply (rule iffI) using nonempty_path_components_of path_component_of_eq_empty apply fastforce by (simp add: path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" apply (simp) by (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] apply safe using assms homeomorphic_map_path_component_of apply fastforce by (metis assms homeomorphic_map_path_component_of imageI) subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp) next case equal then show ?thesis by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" and 2: "2 \ DIM('a)" shows "path_connected (- s)" proof (cases "s = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ s" by auto { fix x y assume "x \ s" "y \ s" then have "x \ a" "y \ a" using \a \ s\ by auto then have bxy: "bounded(insert x (insert y s))" by (simp add: \bounded s\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec) using \x \ a\ \x \ s\ \0 \ u\ CC apply (auto simp: xeq) done } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec) using \y \ a\ \y \ s\ \0 \ u\ DD apply (auto simp: xeq) done } then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" apply (rule path_component_of_subset [of "sphere a B"]) using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) using \x \ a\ using \y \ a\ B apply (auto simp: dist_norm C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" "2 \ DIM('a)" shows "connected (- s)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes s :: "'a :: euclidean_space set" assumes "connected s" "cball a r \ s" "2 \ DIM('a)" shows "connected (s - ball a r)" apply (rule connected_diff_open_from_closed [OF ball_subset_cball]) using assms connected_sphere apply (auto simp: cball_diff_eq_sphere dist_norm) done proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast have "dist a (a + \ *\<^sub>R (SOME i. i \ Basis)) = \" by (simp add: dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "\{S - ball a r |r. 0 < r \ r < \} \ {} \ False" apply (drule_tac c="a + scaleR (\) ((SOME i. i \ Basis))" in subsetD) by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x apply (rule UnionI [of "S - ball a (min \ (dist a x) / 2)"]) using that \0 < \\ apply simp_all apply (rule_tac x="min \ (dist a x) / 2" in exI) apply auto done then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis apply (simp add: abs_if split: if_split_asm) apply (metis add.inverse_inverse real_vector.scale_minus_left xb) done qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" apply (rule connected_intermediate_closure [of "ball a r"]) using assms by auto moreover have "connected {x. r \ dist a x \ x \ S}" apply (rule connected_intermediate_closure [of "- cball a r"]) using assms apply (auto intro: connected_complement_bounded_convex) apply (metis ComplI interior_cball interior_closure mem_ball not_less) done ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) apply (rule_tac x=0 in exI, simp) using vector_choose_size zero_le_one by blast next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis apply (subst *) apply (rule path_connected_continuous_image, auto) done qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes s :: "'a::real_normed_vector set" shows "open s \ open (connected_component_set s x)" apply (simp add: open_contains_ball, clarify) apply (rename_tac y) apply (drule_tac x=y in bspec) apply (simp add: connected_component_in, clarify) apply (rule_tac x=e in exI) by (metis mem_Collect_eq connected_component_eq connected_component_maximal centre_in_ball connected_ball) corollary open_components: fixes s :: "'a::real_normed_vector set" shows "\open u; s \ components u\ \ open s" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes s :: "'a::real_normed_vector set" assumes x: "x \ s" and s: "open s" shows "x \ closure (connected_component_set s y) \ x \ connected_component_set s y" proof - { assume "x \ closure (connected_component_set s y)" moreover have "x \ connected_component_set s x" using x by simp ultimately have "x \ connected_component_set s y" using s by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto - from connectedD [OF \connected S\ 1 2 3 4] + from connectedD [OF \connected S\ 1 2 4 3] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" apply (rule connected_disjoint_Union_open_pick [OF B, of A]) using SA SB \S \ A\ by auto moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" apply (rule connected_disjoint_Union_open_pick [OF A, of B]) using SA SB \T \ B\ by auto ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis apply (rule connected_disjoint_Union_open_unique) apply (simp add: components_eq disjnt_def pairwise_def) using \open S\ apply (simp_all add: assms open_components in_components_connected in_components_nonempty) done qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes "bounded (-s)" shows "\x. x \ s \ \ bounded (connected_component_set s x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" apply (auto simp: bounded_def dist_norm) apply (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) apply simp using i apply (auto simp: algebra_simps) done have **: "{x. B \ i \ x} \ connected_component_set s (B *\<^sub>R i)" apply (rule connected_component_maximal) apply (auto simp: i intro: convex_connected convex_halfspace_ge [of B]) apply (rule *) apply (rule order_trans [OF _ Basis_le_norm [OF i]]) by (simp add: inner_commute) have "B *\<^sub>R i \ s" by (rule *) (simp add: norm_Basis [OF i]) then show ?thesis apply (rule_tac x="B *\<^sub>R i" in exI, clarify) apply (frule bounded_subset [of _ "{x. B \ i \ x}", OF _ **]) using unbounded_inner apply blast done qed lemma cobounded_unique_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes bs: "bounded (-s)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set s x)" "\ bounded(connected_component_set s y)" shows "connected_component_set s x = connected_component_set s y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have ccb: "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) obtain x' where x': "connected_component s x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component s y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component s x' y'" apply (simp add: connected_component_def) apply (rule_tac x="- ball 0 B" in exI) using x' y' apply (auto simp: ccb dist_norm *) done show ?thesis apply (rule connected_component_eq) using x' y' x'y' by (metis (no_types, lifting) connected_component_eq_empty connected_component_eq_eq connected_component_idemp connected_component_in) qed lemma cobounded_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "bounded (-s) \ \c. c \ components s \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "\bounded (- s); c \ components s; \ bounded c; c' \ components s; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac s="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) apply clarify apply (metis connected_component_eq_eq connected_component_in) done lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt) apply (simp add: Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" apply (rule connected_componentI [OF _ subset_refl]) apply (rule connected_complement_bounded_convex) using assms yz by (auto simp: dist_norm) then have "connected_component (- S) y z" apply (rule connected_component_of_subset) apply (metis Bno Compl_anti_mono mem_cball_0 subset_iff) done } note cyz = this show ?thesis apply (auto simp: outside) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) apply (simp add: bounded_pos) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" apply (rule connected_component_maximal) using that interior_subset mem_ball apply auto done then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "interior s \ inside (frontier s)" proof - { fix x y assume x: "x \ interior s" and y: "y \ s" and cc: "connected_component (- frontier s) x y" have "connected_component_set (- frontier s) x \ frontier s \ {}" apply (rule connected_Int_frontier, simp) apply (metis IntI cc connected_component_in connected_component_refl empty_iff interiorE mem_Collect_eq rev_subsetD x) using y cc by blast then have "bounded (connected_component_set (- frontier s) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- s) x y; x \ inside s\ \ y \ inside s" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- s) x y; x \ outside s\ \ y \ outside s" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "convex s" and z: "z \ s" shows "z \ outside s" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain a where "a \ s" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- s) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- s) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ s" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF s \a \ s\ ins, of "1/(u*C + 1)"] \C>0\ \z \ s\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- s) z (z + C *\<^sub>R (z-a))" apply (rule connected_componentI [OF connected_segment [of z "z + C *\<^sub>R (z-a)"]]) apply (simp add: closed_segment_def) using contra apply auto done then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex s" shows "outside s = - s" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "convex s \ inside s = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex t; s \ t\ \ - t \ outside s" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "outside(frontier s) \ - closure s" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior s \ inside (frontier s)" hence "interior s \ inside (frontier s) = inside (frontier s)" by (simp add: subset_Un_eq) then have "closure s \ frontier s \ inside (frontier s)" using frontier_def by auto } then show "closure s \ frontier s \ inside (frontier s)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" "convex s" shows "outside(frontier s) = - closure s" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded s; convex s\ \ inside(frontier s) = interior s" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (inside s)" proof - { fix x assume x: "x \ inside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside s" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (outside s)" proof - { fix x assume x: "x \ outside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis Int_iff outside_def connected_component_refl_eq x) then have "\e>0. ball x e \ outside s" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(inside s) \ s \ inside s" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(inside s) \ s" proof - have "closure (inside s) \ - inside s = closure (inside s) - interior (inside s)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside s \ - outside s = s" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside s) \ - outside s" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside s) - interior (inside s) \ s" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(outside s) \ s \ outside s" apply (rule closure_minimal, simp) by (metis assms closed_open inside_outside open_inside) lemma frontier_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(outside s) \ s" apply (simp add: frontier_def open_outside interior_open) by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup.commute) lemma inside_complement_unbounded_connected_empty: "\connected (- s); \ bounded (- s)\ \ inside s = {}" apply (simp add: inside_def) by (meson Compl_iff bounded_subset connected_component_maximal order_refl) lemma inside_bounded_complement_connected_empty: fixes s :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- s); bounded s\ \ inside s = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "s \ inside t" shows "inside s - t \ inside t" unfolding inside_def proof clarify fix x assume x: "x \ t" "x \ s" and bo: "bounded (connected_component_set (- s) x)" show "bounded (connected_component_set (- t) x)" proof (cases "s \ connected_component_set (- t) x = {}") case True show ?thesis apply (rule bounded_subset [OF bo]) apply (rule connected_component_maximal) using x True apply auto done next case False then show ?thesis using assms [unfolded inside_def] x apply (simp add: disjoint_iff_not_equal, clarify) apply (drule subsetD, assumption, auto) by (metis (no_types, hide_lams) ComplI connected_component_eq_eq) qed qed lemma inside_inside_subset: "inside(inside s) \ s" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected t; inside s \ t \ {}; outside s \ t \ {}\ \ s \ t \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" shows "outside s \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes s :: "'a :: {real_normed_vector,perfect_space} set" assumes s: "compact s" and t: "open t" and "s \ t" "t \ {}" shows "outside s \ t \ {}" proof - have "outside s \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty s) with assms obtain a b where a: "a \ outside s" and b: "b \ t" by auto show ?thesis proof (cases "a \ t") case True with a show ?thesis by blast next case False have front: "frontier t \ - s" using \s \ t\ frontier_disjoint_eq t by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -s" unfolding c_def using front pf by blast moreover have "open (-s)" using s compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -s" using open_contains_cball[of "-s"] s by blast then obtain d where "d \ t" and d: "dist d c < \" using closure_approachable [of c t] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -s" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -s" using pimg_sbs apply (auto simp: path_image_def) apply (drule subsetD) using \c \ - s\ \s \ t\ interior_subset apply (auto simp: c_def) done have "closed_segment c d \ cball c \" apply (simp add: segment_convex_hull) apply (rule hull_minimal) using \\ > 0\ d apply (auto simp: dist_commute) done with \ have "closed_segment c d \ -s" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- s) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside s \ t \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ t\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- t)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes s :: "'a :: euclidean_space set" assumes s: "closed s" and t: "compact t" and "connected t" "s \ inside t" shows "inside s \ inside t" proof (cases "inside t = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have coms: "compact s" using assms apply (simp add: s compact_eq_bounded_closed) by (meson bounded_inside bounded_subset compact_imp_bounded) then have bst: "bounded (s \ t)" by (simp add: compact_imp_bounded t) then obtain r where "0 < r" and r: "s \ t \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside s \ outside t \ {}" proof - have "- ball 0 r \ outside s" apply (rule outside_subset_convex) using r by auto moreover have "- ball 0 r \ outside t" apply (rule outside_subset_convex) using r by auto ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "s \ t = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside s \ inside t \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t) ultimately have "inside s \ t = {}" using inside_outside_intersect_connected [OF \connected t\, of s] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \s \ inside t\] by blast qed qed lemma connected_with_inside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ inside s)" proof (cases "s \ inside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ inside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ inside s)" if "a \ (s \ inside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ inside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_inside_subset s apply fastforce by (metis (no_types, lifting) frontier_inside_subset insertE insert_Diff interior_eq open_inside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ outside s)" proof (cases "s \ outside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ outside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ outside s)" if "a \ (s \ outside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ outside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_outside_subset s apply fastforce by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "closed s" and cons: "connected s" shows "inside (inside s) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside s \ components (- s) \ connected(inside s) \ inside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: inside_same_component connected_componentI) apply (metis IntI empty_iff inside_no_overlap) done text\The proof is virtually the same as that above.\ lemma outside_in_components: "outside s \ components (- s) \ connected(outside s) \ outside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: outside_same_component connected_componentI) apply (metis IntI empty_iff outside_no_overlap) done lemma bounded_unique_outside: fixes s :: "'a :: euclidean_space set" shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ \ bounded c \ c = outside s)" apply (rule iffI) apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" apply (rule distance_attains_inf [of "frontier(f ` S)" "f a"]) using \a \ S\ by (auto simp: frontier_eq_empty dist_norm False) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ apply (simp add: compact_closure [symmetric] compact_def) apply (drule_tac x=z in spec) using closure_subset apply force done then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" apply (rule le_no) using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast show ?thesis apply (rule **) (*such a horrible mess*) apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball]) using \a \ S\ \0 < r\ apply (auto simp: disjoint_iff_not_equal dist_norm) by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed subsubsection\Special characterizations of classes of functions into and out of R.\ lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" proof - have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" if "x \ y" for x y :: 'a proof (intro exI conjI) let ?r = "dist x y / 2" have [simp]: "?r > 0" by (simp add: that) show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" by (auto simp add: that) show "disjnt (ball x ?r) (ball y ?r)" unfolding disjnt_def by (simp add: disjoint_ballI) qed then show ?thesis by (simp add: Hausdorff_space_def) qed proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast subsection\<^marker>\tag unimportant\ \Rectangular paths\ definition\<^marker>\tag unimportant\ rectpath where "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" lemma path_rectpath [simp, intro]: "path (rectpath a b)" by (simp add: Let_def rectpath_def) lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma simple_path_rectpath [simp, intro]: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "simple_path (rectpath a1 a3)" unfolding rectpath_def Let_def using assms by (intro simple_path_join_loop arc_join arc_linepath) (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) lemma path_image_rectpath: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "path_image (rectpath a1 a3) = {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") proof - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ closed_segment a4 a3 \ closed_segment a1 a4" by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute a2_def a4_def Un_assoc) also have "\ = ?rhs" using assms by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) finally show ?thesis . qed lemma path_image_rectpath_subset_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ cbox a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) lemma path_image_rectpath_inter_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ box a b = {}" using assms by (auto simp: path_image_rectpath in_box_complex_iff) lemma path_image_rectpath_cbox_minus_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) = cbox a b - box a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff) end diff --git a/src/HOL/Analysis/Polytope.thy b/src/HOL/Analysis/Polytope.thy --- a/src/HOL/Analysis/Polytope.thy +++ b/src/HOL/Analysis/Polytope.thy @@ -1,4005 +1,4005 @@ section \Faces, Extreme Points, Polytopes, Polyhedra etc\ text\Ported from HOL Light by L C Paulson\ theory Polytope imports Cartesian_Euclidean_Space Path_Connected begin subsection \Faces of a (usually convex) set\ definition\<^marker>\tag important\ face_of :: "['a::real_vector set, 'a set] \ bool" (infixr "(face'_of)" 50) where "T face_of S \ T \ S \ convex T \ (\a \ S. \b \ S. \x \ T. x \ open_segment a b \ a \ T \ b \ T)" lemma face_ofD: "\T face_of S; x \ open_segment a b; a \ S; b \ S; x \ T\ \ a \ T \ b \ T" unfolding face_of_def by blast lemma face_of_translation_eq [simp]: "((+) a ` T face_of (+) a ` S) \ T face_of S" proof - have *: "\a T S. T face_of S \ ((+) a ` T face_of (+) a ` S)" apply (simp add: face_of_def Ball_def, clarify) by (meson imageI open_segment_translation_eq) show ?thesis apply (rule iffI) apply (force simp: image_comp o_def dest: * [where a = "-a"]) apply (blast intro: *) done qed lemma face_of_linear_image: assumes "linear f" "inj f" shows "(f ` c face_of f ` S) \ c face_of S" by (simp add: face_of_def inj_image_subset_iff inj_image_mem_iff open_segment_linear_image assms) lemma face_of_refl: "convex S \ S face_of S" by (auto simp: face_of_def) lemma face_of_refl_eq: "S face_of S \ convex S" by (auto simp: face_of_def) lemma empty_face_of [iff]: "{} face_of S" by (simp add: face_of_def) lemma face_of_empty [simp]: "S face_of {} \ S = {}" by (meson empty_face_of face_of_def subset_empty) lemma face_of_trans [trans]: "\S face_of T; T face_of u\ \ S face_of u" unfolding face_of_def by (safe; blast) lemma face_of_face: "T face_of S \ (f face_of T \ f face_of S \ f \ T)" unfolding face_of_def by (safe; blast) lemma face_of_subset: "\F face_of S; F \ T; T \ S\ \ F face_of T" unfolding face_of_def by (safe; blast) lemma face_of_slice: "\F face_of S; convex T\ \ (F \ T) face_of (S \ T)" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Int: "\t1 face_of S; t2 face_of S\ \ (t1 \ t2) face_of S" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_Inter: "\A \ {}; \T. T \ A \ T face_of S\ \ (\ A) face_of S" unfolding face_of_def by (blast intro: convex_Inter) lemma face_of_Int_Int: "\F face_of T; F' face_of t'\ \ (F \ F') face_of (T \ t')" unfolding face_of_def by (blast intro: convex_Int) lemma face_of_imp_subset: "T face_of S \ T \ S" unfolding face_of_def by blast proposition face_of_imp_eq_affine_Int: fixes S :: "'a::euclidean_space set" assumes S: "convex S" and T: "T face_of S" shows "T = (affine hull T) \ S" proof - have "convex T" using T by (simp add: face_of_def) have *: False if x: "x \ affine hull T" and "x \ S" "x \ T" and y: "y \ rel_interior T" for x y proof - obtain e where "e>0" and e: "cball y e \ affine hull T \ T" using y by (auto simp: rel_interior_cball) have "y \ x" "y \ S" "y \ T" using face_of_imp_subset rel_interior_subset T that by blast+ then have zne: "\u. \u \ {0<..<1}; (1 - u) *\<^sub>R y + u *\<^sub>R x \ T\ \ False" using \x \ S\ \x \ T\ \T face_of S\ unfolding face_of_def apply clarify apply (drule_tac x=x in bspec, assumption) apply (drule_tac x=y in bspec, assumption) apply (subst (asm) open_segment_commute) apply (force simp: open_segment_image_interval image_def) done have in01: "min (1/2) (e / norm (x - y)) \ {0<..<1}" using \y \ x\ \e > 0\ by simp show ?thesis apply (rule zne [OF in01]) apply (rule e [THEN subsetD]) apply (rule IntI) using \y \ x\ \e > 0\ apply (simp add: cball_def dist_norm algebra_simps) apply (simp add: Real_Vector_Spaces.scaleR_diff_right [symmetric] norm_minus_commute min_mult_distrib_right) apply (rule mem_affine [OF affine_affine_hull _ x]) using \y \ T\ apply (auto simp: hull_inc) done qed show ?thesis apply (rule subset_antisym) using assms apply (simp add: hull_subset face_of_imp_subset) apply (cases "T={}", simp) apply (force simp: rel_interior_eq_empty [symmetric] \convex T\ intro: *) done qed lemma face_of_imp_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "T face_of S" shows "closed T" by (metis affine_affine_hull affine_closed closed_Int face_of_imp_eq_affine_Int assms) lemma face_of_Int_supporting_hyperplane_le_strong: assumes "convex(S \ {x. a \ x = b})" and aleb: "\x. x \ S \ a \ x \ b" shows "(S \ {x. a \ x = b}) face_of S" proof - have *: "a \ u = a \ x" if "x \ open_segment u v" "u \ S" "v \ S" and b: "b = a \ x" for u v x proof (rule antisym) show "a \ u \ a \ x" using aleb \u \ S\ \b = a \ x\ by blast next obtain \ where "b = a \ ((1 - \) *\<^sub>R u + \ *\<^sub>R v)" "0 < \" "\ < 1" using \b = a \ x\ \x \ open_segment u v\ in_segment by (auto simp: open_segment_image_interval split: if_split_asm) then have "b + \ * (a \ u) \ a \ u + \ * b" using aleb [OF \v \ S\] by (simp add: algebra_simps) then have "(1 - \) * b \ (1 - \) * (a \ u)" by (simp add: algebra_simps) then have "b \ a \ u" using \\ < 1\ by auto with b show "a \ x \ a \ u" by simp qed show ?thesis apply (simp add: face_of_def assms) using "*" open_segment_commute by blast qed lemma face_of_Int_supporting_hyperplane_ge_strong: "\convex(S \ {x. a \ x = b}); \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" using face_of_Int_supporting_hyperplane_le_strong [of S "-a" "-b"] by simp lemma face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_le_strong) lemma face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) face_of S" by (simp add: convex_Int convex_hyperplane face_of_Int_supporting_hyperplane_ge_strong) lemma face_of_imp_convex: "T face_of S \ convex T" using face_of_def by blast lemma face_of_imp_compact: fixes S :: "'a::euclidean_space set" shows "\convex S; compact S; T face_of S\ \ compact T" by (meson bounded_subset compact_eq_bounded_closed face_of_imp_closed face_of_imp_subset) lemma face_of_Int_subface: "\A \ B face_of A; A \ B face_of B; C face_of A; D face_of B\ \ (C \ D) face_of C \ (C \ D) face_of D" by (meson face_of_Int_Int face_of_face inf_le1 inf_le2) lemma subset_of_face_of: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u \ S" "T \ (rel_interior u) \ {}" shows "u \ T" proof fix c assume "c \ u" obtain b where "b \ T" "b \ rel_interior u" using assms by auto then obtain e where "e>0" "b \ u" and e: "cball b e \ affine hull u \ u" by (auto simp: rel_interior_cball) show "c \ T" proof (cases "b=c") case True with \b \ T\ show ?thesis by blast next case False define d where "d = b + (e / norm(b - c)) *\<^sub>R (b - c)" have "d \ cball b e \ affine hull u" using \e > 0\ \b \ u\ \c \ u\ by (simp add: d_def dist_norm hull_inc mem_affine_3_minus False) with e have "d \ u" by blast have nbc: "norm (b - c) + e > 0" using \e > 0\ by (metis add.commute le_less_trans less_add_same_cancel2 norm_ge_zero) then have [simp]: "d \ c" using False scaleR_cancel_left [of "1 + (e / norm (b - c))" b c] by (simp add: algebra_simps d_def) (simp add: field_split_simps) have [simp]: "((e - e * e / (e + norm (b - c))) / norm (b - c)) = (e / (e + norm (b - c)))" using False nbc by (simp add: divide_simps) (simp add: algebra_simps) have "b \ open_segment d c" apply (simp add: open_segment_image_interval) apply (simp add: d_def algebra_simps image_def) apply (rule_tac x="e / (e + norm (b - c))" in bexI) using False nbc \0 < e\ apply (auto simp: algebra_simps) done then have "d \ T \ c \ T" apply (rule face_ofD [OF \T face_of S\]) using \d \ u\ \c \ u\ \u \ S\ \b \ T\ apply auto done then show ?thesis .. qed qed lemma face_of_eq: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "u face_of S" "(rel_interior T) \ (rel_interior u) \ {}" shows "T = u" apply (rule subset_antisym) apply (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subsetCE subset_of_face_of) by (metis assms disjoint_iff_not_equal face_of_imp_subset rel_interior_subset subset_iff subset_of_face_of) lemma face_of_disjoint_rel_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_interior S = {}" by (meson assms subset_of_face_of face_of_imp_subset order_refl subset_antisym) lemma face_of_disjoint_interior: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ interior S = {}" proof - have "T \ interior S \ rel_interior S" by (meson inf_sup_ord(2) interior_subset_rel_interior order.trans) thus ?thesis by (metis (no_types) Int_greatest assms face_of_disjoint_rel_interior inf_sup_ord(1) subset_empty) qed lemma face_of_subset_rel_boundary: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ (S - rel_interior S)" by (meson DiffI assms disjoint_iff_not_equal face_of_disjoint_rel_interior face_of_imp_subset rev_subsetD subsetI) lemma face_of_subset_rel_frontier: fixes S :: "'a::real_normed_vector set" assumes "T face_of S" "T \ S" shows "T \ rel_frontier S" using assms closure_subset face_of_disjoint_rel_interior face_of_imp_subset rel_frontier_def by fastforce lemma face_of_aff_dim_lt: fixes S :: "'a::euclidean_space set" assumes "convex S" "T face_of S" "T \ S" shows "aff_dim T < aff_dim S" proof - have "aff_dim T \ aff_dim S" by (simp add: face_of_imp_subset aff_dim_subset assms) moreover have "aff_dim T \ aff_dim S" proof (cases "T = {}") case True then show ?thesis by (metis aff_dim_empty \T \ S\) next case False then show ?thesis by (metis Set.set_insert assms convex_rel_frontier_aff_dim dual_order.irrefl face_of_imp_convex face_of_subset_rel_frontier insert_not_empty subsetI) qed ultimately show ?thesis by simp qed lemma subset_of_face_of_affine_hull: fixes S :: "'a::euclidean_space set" assumes T: "T face_of S" and "convex S" "U \ S" and dis: "\ disjnt (affine hull T) (rel_interior U)" shows "U \ T" apply (rule subset_of_face_of [OF T \U \ S\]) using face_of_imp_eq_affine_Int [OF \convex S\ T] using rel_interior_subset [of U] dis using \U \ S\ disjnt_def by fastforce lemma affine_hull_face_of_disjoint_rel_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" "F face_of S" "F \ S" shows "affine hull F \ rel_interior S = {}" by (metis assms disjnt_def face_of_imp_subset order_refl subset_antisym subset_of_face_of_affine_hull) lemma affine_diff_divide: assumes "affine S" "k \ 0" "k \ 1" and xy: "x \ S" "y /\<^sub>R (1 - k) \ S" shows "(x - y) /\<^sub>R k \ S" proof - have "inverse(k) *\<^sub>R (x - y) = (1 - inverse k) *\<^sub>R inverse(1 - k) *\<^sub>R y + inverse(k) *\<^sub>R x" using assms by (simp add: algebra_simps) (simp add: scaleR_left_distrib [symmetric] field_split_simps) then show ?thesis using \affine S\ xy by (auto simp: affine_alt) qed proposition face_of_convex_hulls: assumes S: "finite S" "T \ S" and disj: "affine hull T \ convex hull (S - T) = {}" shows "(convex hull T) face_of (convex hull S)" proof - have fin: "finite T" "finite (S - T)" using assms by (auto simp: finite_subset) have *: "x \ convex hull T" if x: "x \ convex hull S" and y: "y \ convex hull S" and w: "w \ convex hull T" "w \ open_segment x y" for x y w proof - have waff: "w \ affine hull T" using convex_hull_subset_affine_hull w by blast obtain a b where a: "\i. i \ S \ 0 \ a i" and asum: "sum a S = 1" and aeqx: "(\i\S. a i *\<^sub>R i) = x" and b: "\i. i \ S \ 0 \ b i" and bsum: "sum b S = 1" and beqy: "(\i\S. b i *\<^sub>R i) = y" using x y by (auto simp: assms convex_hull_finite) obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \ convex hull T" "x \ y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y" and u01: "0 < u" "u < 1" using w by (auto simp: open_segment_image_interval split: if_split_asm) define c where "c i = (1 - u) * a i + u * b i" for i have cge0: "\i. i \ S \ 0 \ c i" using a b u01 by (simp add: c_def) have sumc1: "sum c S = 1" by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum) have sumci_xy: "(\i\S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y" apply (simp add: c_def sum.distrib scaleR_left_distrib) by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy) show ?thesis proof (cases "sum c (S - T) = 0") case True have ci0: "\i. i \ (S - T) \ c i = 0" using True cge0 fin(2) sum_nonneg_eq_0_iff by auto have a0: "a i = 0" if "i \ (S - T)" for i using ci0 [OF that] u01 a [of i] b [of i] that by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff) have [simp]: "sum a T = 1" using assms by (metis sum.mono_neutral_cong_right a0 asum) show ?thesis apply (simp add: convex_hull_finite \finite T\) apply (rule_tac x=a in exI) using a0 assms apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right) done next case False define k where "k = sum c (S - T)" have "k > 0" using False unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less) have weq_sumsum: "w = sum (\x. c x *\<^sub>R x) T + sum (\x. c x *\<^sub>R x) (S - T)" by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq) show ?thesis proof (cases "k = 1") case True then have "sum c T = 0" by (simp add: S k_def sum_diff sumc1) then have [simp]: "sum c (S - T) = 1" by (simp add: S sum_diff sumc1) have ci0: "\i. i \ T \ c i = 0" by (meson \finite T\ \sum c T = 0\ \T \ S\ cge0 sum_nonneg_eq_0_iff subsetCE) then have [simp]: "(\i\S-T. c i *\<^sub>R i) = w" by (simp add: weq_sumsum) have "w \ convex hull (S - T)" apply (simp add: convex_hull_finite fin) apply (rule_tac x=c in exI) apply (auto simp: cge0 weq True k_def) done then show ?thesis using disj waff by blast next case False then have sumcf: "sum c T = 1 - k" by (simp add: S k_def sum_diff sumc1) have "(\i\T. c i *\<^sub>R i) /\<^sub>R (1 - k) \ convex hull T" apply (simp add: convex_hull_finite fin) apply (rule_tac x="\i. inverse (1-k) * c i" in exI) apply auto apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE) apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf) by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong) with \0 < k\ have "inverse(k) *\<^sub>R (w - sum (\i. c i *\<^sub>R i) T) \ affine hull T" by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD]) moreover have "inverse(k) *\<^sub>R (w - sum (\x. c x *\<^sub>R x) T) \ convex hull (S - T)" apply (simp add: weq_sumsum convex_hull_finite fin) apply (rule_tac x="\i. inverse k * c i" in exI) using \k > 0\ cge0 apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric]) done ultimately show ?thesis using disj by blast qed qed qed have [simp]: "convex hull T \ convex hull S" by (simp add: \T \ S\ hull_mono) show ?thesis using open_segment_commute by (auto simp: face_of_def intro: *) qed proposition face_of_convex_hull_insert: "\finite S; a \ affine hull S; T face_of convex hull S\ \ T face_of convex hull insert a S" apply (rule face_of_trans, blast) apply (rule face_of_convex_hulls; force simp: insert_Diff_if) done proposition face_of_affine_trivial: assumes "affine S" "T face_of S" shows "T = {} \ T = S" proof (rule ccontr, clarsimp) assume "T \ {}" "T \ S" then obtain a where "a \ T" by auto then have "a \ S" using \T face_of S\ face_of_imp_subset by blast have "S \ T" proof fix b assume "b \ S" show "b \ T" proof (cases "a = b") case True with \a \ T\ show ?thesis by auto next case False then have "a \ open_segment (2 *\<^sub>R a - b) b" apply (auto simp: open_segment_def closed_segment_def) apply (rule_tac x="1/2" in exI) apply (simp add: algebra_simps) by (simp add: scaleR_2) moreover have "2 *\<^sub>R a - b \ S" by (rule mem_affine [OF \affine S\ \a \ S\ \b \ S\, of 2 "-1", simplified]) moreover note \b \ S\ \a \ T\ ultimately show ?thesis by (rule face_ofD [OF \T face_of S\, THEN conjunct2]) qed qed then show False using \T \ S\ \T face_of S\ face_of_imp_subset by blast qed lemma face_of_affine_eq: "affine S \ (T face_of S \ T = {} \ T = S)" using affine_imp_convex face_of_affine_trivial face_of_refl by auto proposition Inter_faces_finite_altbound: fixes T :: "'a::euclidean_space set set" assumes cfaI: "\c. c \ T \ c face_of S" shows "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ \F' = \T" proof (cases "\F'. finite F' \ F' \ T \ card F' \ DIM('a) + 2 \ (\c. c \ T \ c \ (\F') \ (\F'))") case True then obtain c where c: "\F'. \finite F'; F' \ T; card F' \ DIM('a) + 2\ \ c F' \ T \ c F' \ (\F') \ (\F')" by metis define d where "d = rec_nat {c{}} (\n r. insert (c r) r)" have [simp]: "d 0 = {c {}}" by (simp add: d_def) have dSuc [simp]: "\n. d (Suc n) = insert (c (d n)) (d n)" by (simp add: d_def) have dn_notempty: "d n \ {}" for n by (induction n) auto have dn_le_Suc: "d n \ T \ finite(d n) \ card(d n) \ Suc n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: c) next case (Suc n) then show ?case by (auto simp: c card_insert_if) qed have aff_dim_le: "aff_dim(\(d n)) \ DIM('a) - int n" if "n \ DIM('a) + 2" for n using that proof (induction n) case 0 then show ?case by (simp add: aff_dim_le_DIM) next case (Suc n) have fs: "\(d (Suc n)) face_of S" by (meson Suc.prems cfaI dn_le_Suc dn_notempty face_of_Inter subsetCE) have condn: "convex (\(d n))" using Suc.prems nat_le_linear not_less_eq_eq by (blast intro: face_of_imp_convex cfaI convex_Inter dest: dn_le_Suc) have fdn: "\(d (Suc n)) face_of \(d n)" by (metis (no_types, lifting) Inter_anti_mono Suc.prems dSuc cfaI dn_le_Suc dn_notempty face_of_Inter face_of_imp_subset face_of_subset subset_iff subset_insertI) have ne: "\(d (Suc n)) \ \(d n)" by (metis (no_types, lifting) Suc.prems Suc_leD c complete_lattice_class.Inf_insert dSuc dn_le_Suc less_irrefl order.trans) have *: "\m::int. \d. \d'::int. d < d' \ d' \ m - n \ d \ m - of_nat(n+1)" by arith have "aff_dim (\(d (Suc n))) < aff_dim (\(d n))" by (rule face_of_aff_dim_lt [OF condn fdn ne]) moreover have "aff_dim (\(d n)) \ int (DIM('a)) - int n" using Suc by auto ultimately have "aff_dim (\(d (Suc n))) \ int (DIM('a)) - (n+1)" by arith then show ?case by linarith qed have "aff_dim (\(d (DIM('a) + 2))) \ -2" using aff_dim_le [OF order_refl] by simp with aff_dim_geq [of "\(d (DIM('a) + 2))"] show ?thesis using order.trans by fastforce next case False then show ?thesis apply simp apply (erule ex_forward) by blast qed lemma faces_of_translation: "{F. F face_of image (\x. a + x) S} = image (image (\x. a + x)) {F. F face_of S}" apply (rule subset_antisym, clarify) apply (auto simp: image_iff) apply (metis face_of_imp_subset face_of_translation_eq subset_imageE) done proposition face_of_Times: assumes "F face_of S" and "F' face_of S'" shows "(F \ F') face_of (S \ S')" proof - have "F \ F' \ S \ S'" using assms [unfolded face_of_def] by blast moreover have "convex (F \ F')" using assms [unfolded face_of_def] by (blast intro: convex_Times) moreover have "a \ F \ a' \ F' \ b \ F \ b' \ F'" if "a \ S" "b \ S" "a' \ S'" "b' \ S'" "x \ F \ F'" "x \ open_segment (a,a') (b,b')" for a b a' b' x proof (cases "b=a \ b'=a'") case True with that show ?thesis using assms by (force simp: in_segment dest: face_ofD) next case False with assms [unfolded face_of_def] that show ?thesis by (blast dest!: open_segment_PairD) qed ultimately show ?thesis unfolding face_of_def by blast qed corollary face_of_Times_decomp: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "c face_of (S \ S') \ (\F F'. F face_of S \ F' face_of S' \ c = F \ F')" (is "?lhs = ?rhs") proof assume c: ?lhs show ?rhs proof (cases "c = {}") case True then show ?thesis by auto next case False have 1: "fst ` c \ S" "snd ` c \ S'" using c face_of_imp_subset by fastforce+ have "convex c" using c by (metis face_of_imp_convex) have conv: "convex (fst ` c)" "convex (snd ` c)" - by (simp_all add: \convex c\ convex_linear_image fst_linear snd_linear) + by (simp_all add: \convex c\ convex_linear_image linear_fst linear_snd) have fstab: "a \ fst ` c \ b \ fst ` c" if "a \ S" "b \ S" "x \ open_segment a b" "(x,x') \ c" for a b x x' proof - have *: "(x,x') \ open_segment (a,x') (b,x')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF c *] that face_of_imp_subset [OF c] by force qed have fst: "fst ` c face_of S" by (force simp: face_of_def 1 conv fstab) have sndab: "a' \ snd ` c \ b' \ snd ` c" if "a' \ S'" "b' \ S'" "x' \ open_segment a' b'" "(x,x') \ c" for a' b' x x' proof - have *: "(x,x') \ open_segment (x,a') (x,b')" using that by (auto simp: in_segment) show ?thesis using face_ofD [OF c *] that face_of_imp_subset [OF c] by force qed have snd: "snd ` c face_of S'" by (force simp: face_of_def 1 conv sndab) have cc: "rel_interior c \ rel_interior (fst ` c) \ rel_interior (snd ` c)" - by (force simp: face_of_Times rel_interior_Times conv fst snd \convex c\ fst_linear snd_linear rel_interior_convex_linear_image [symmetric]) + by (force simp: face_of_Times rel_interior_Times conv fst snd \convex c\ linear_fst linear_snd rel_interior_convex_linear_image [symmetric]) have "c = fst ` c \ snd ` c" apply (rule face_of_eq [OF c]) apply (simp_all add: face_of_Times rel_interior_Times conv fst snd) using False rel_interior_eq_empty \convex c\ cc apply blast done with fst snd show ?thesis by metis qed next assume ?rhs with face_of_Times show ?lhs by auto qed lemma face_of_Times_eq: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" shows "(F \ F') face_of (S \ S') \ F = {} \ F' = {} \ F face_of S \ F' face_of S'" by (auto simp: face_of_Times_decomp times_eq_iff) lemma hyperplane_face_of_halfspace_le: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_le [OF convex_halfspace_le [of a b], of a b] show ?thesis by auto qed lemma hyperplane_face_of_halfspace_ge: "{x. a \ x = b} face_of {x. a \ x \ b}" proof - have "{x. a \ x \ b} \ {x. a \ x = b} = {x. a \ x = b}" by auto with face_of_Int_supporting_hyperplane_ge [OF convex_halfspace_ge [of b a], of b a] show ?thesis by auto qed lemma face_of_halfspace_le: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis using face_of_affine_eq affine_UNIV by auto next case False then have ine: "interior {x. a \ x \ b} \ {}" using halfspace_eq_empty_lt interior_halfspace_le by blast show ?thesis proof assume L: ?lhs have "F \ {x. a \ x \ b} \ F face_of {x. a \ x = b}" using False apply (simp add: frontier_halfspace_le [symmetric] rel_frontier_nonempty_interior [OF ine, symmetric]) apply (rule face_of_subset [OF L]) apply (simp add: face_of_subset_rel_frontier [OF L]) apply (force simp: rel_frontier_def closed_halfspace_le) done with L show ?rhs using affine_hyperplane face_of_affine_eq by blast next assume ?rhs then show ?lhs by (metis convex_halfspace_le empty_face_of face_of_refl hyperplane_face_of_halfspace_le) qed qed lemma face_of_halfspace_ge: fixes a :: "'n::euclidean_space" shows "F face_of {x. a \ x \ b} \ F = {} \ F = {x. a \ x = b} \ F = {x. a \ x \ b}" using face_of_halfspace_le [of F "-a" "-b"] by simp subsection\Exposed faces\ text\That is, faces that are intersection with supporting hyperplane\ definition\<^marker>\tag important\ exposed_face_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(exposed'_face'_of)" 50) where "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b})" lemma empty_exposed_face_of [iff]: "{} exposed_face_of S" apply (simp add: exposed_face_of_def) apply (rule_tac x=0 in exI) apply (rule_tac x=1 in exI, force) done lemma exposed_face_of_refl_eq [simp]: "S exposed_face_of S \ convex S" apply (simp add: exposed_face_of_def face_of_refl_eq, auto) apply (rule_tac x=0 in exI)+ apply force done lemma exposed_face_of_refl: "convex S \ S exposed_face_of S" by simp lemma exposed_face_of: "T exposed_face_of S \ T face_of S \ (T = {} \ T = S \ (\a b. a \ 0 \ S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b}))" proof (cases "T = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "T = S") case True then show ?thesis by (simp add: face_of_refl_eq) next case False with \T \ {}\ show ?thesis apply (auto simp: exposed_face_of_def) apply (metis inner_zero_left) done qed qed lemma exposed_face_of_Int_supporting_hyperplane_le: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" by (force simp: exposed_face_of_def face_of_Int_supporting_hyperplane_le) lemma exposed_face_of_Int_supporting_hyperplane_ge: "\convex S; \x. x \ S \ a \ x \ b\ \ (S \ {x. a \ x = b}) exposed_face_of S" using exposed_face_of_Int_supporting_hyperplane_le [of S "-a" "-b"] by simp proposition exposed_face_of_Int: assumes "T exposed_face_of S" and "u exposed_face_of S" shows "(T \ u) exposed_face_of S" proof - obtain a b where T: "S \ {x. a \ x = b} face_of S" and S: "S \ {x. a \ x \ b}" and teq: "T = S \ {x. a \ x = b}" using assms by (auto simp: exposed_face_of_def) obtain a' b' where u: "S \ {x. a' \ x = b'} face_of S" and s': "S \ {x. a' \ x \ b'}" and ueq: "u = S \ {x. a' \ x = b'}" using assms by (auto simp: exposed_face_of_def) have tu: "T \ u face_of S" using T teq u ueq by (simp add: face_of_Int) have ss: "S \ {x. (a + a') \ x \ b + b'}" using S s' by (force simp: inner_left_distrib) show ?thesis apply (simp add: exposed_face_of_def tu) apply (rule_tac x="a+a'" in exI) apply (rule_tac x="b+b'" in exI) using S s' apply (fastforce simp: ss inner_left_distrib teq ueq) done qed proposition exposed_face_of_Inter: fixes P :: "'a::euclidean_space set set" assumes "P \ {}" and "\T. T \ P \ T exposed_face_of S" shows "\P exposed_face_of S" proof - obtain Q where "finite Q" and QsubP: "Q \ P" "card Q \ DIM('a) + 2" and IntQ: "\Q = \P" using Inter_faces_finite_altbound [of P S] assms [unfolded exposed_face_of] by force show ?thesis proof (cases "Q = {}") case True then show ?thesis by (metis IntQ Inter_UNIV_conv(2) assms(1) assms(2) ex_in_conv) next case False have "Q \ {T. T exposed_face_of S}" using QsubP assms by blast moreover have "Q \ {T. T exposed_face_of S} \ \Q exposed_face_of S" using \finite Q\ False apply (induction Q rule: finite_induct) using exposed_face_of_Int apply fastforce+ done ultimately show ?thesis by (simp add: IntQ) qed qed proposition exposed_face_of_sums: assumes "convex S" and "convex T" and "F exposed_face_of {x + y | x y. x \ S \ y \ T}" (is "F exposed_face_of ?ST") obtains k l where "k exposed_face_of S" "l exposed_face_of T" "F = {x + y | x y. x \ k \ y \ l}" proof (cases "F = {}") case True then show ?thesis using that by blast next case False show ?thesis proof (cases "F = ?ST") case True then show ?thesis using assms exposed_face_of_refl_eq that by blast next case False obtain p where "p \ F" using \F \ {}\ by blast moreover obtain u z where T: "?ST \ {x. u \ x = z} face_of ?ST" and S: "?ST \ {x. u \ x \ z}" and feq: "F = ?ST \ {x. u \ x = z}" using assms by (auto simp: exposed_face_of_def) ultimately obtain a0 b0 where p: "p = a0 + b0" and "a0 \ S" "b0 \ T" and z: "u \ p = z" by auto have lez: "u \ (x + y) \ z" if "x \ S" "y \ T" for x y using S that by auto have sef: "S \ {x. u \ x = u \ a0} exposed_face_of S" apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex S\]) apply (metis p z add_le_cancel_right inner_right_distrib lez [OF _ \b0 \ T\]) done have tef: "T \ {x. u \ x = u \ b0} exposed_face_of T" apply (rule exposed_face_of_Int_supporting_hyperplane_le [OF \convex T\]) apply (metis p z add.commute add_le_cancel_right inner_right_distrib lez [OF \a0 \ S\]) done have "{x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0} \ F" by (auto simp: feq) (metis inner_right_distrib p z) moreover have "F \ {x + y |x y. x \ S \ u \ x = u \ a0 \ y \ T \ u \ y = u \ b0}" apply (auto simp: feq) apply (rename_tac x y) apply (rule_tac x=x in exI) apply (rule_tac x=y in exI, simp) using z p \a0 \ S\ \b0 \ T\ apply clarify apply (simp add: inner_right_distrib) apply (metis add_le_cancel_right antisym lez [unfolded inner_right_distrib] add.commute) done ultimately have "F = {x + y |x y. x \ S \ {x. u \ x = u \ a0} \ y \ T \ {x. u \ x = u \ b0}}" by blast then show ?thesis by (rule that [OF sef tef]) qed qed proposition exposed_face_of_parallel: "T exposed_face_of S \ T face_of S \ (\a b. S \ {x. a \ x \ b} \ T = S \ {x. a \ x = b} \ (T \ {} \ T \ S \ a \ 0) \ (T \ S \ (\w \ affine hull S. (w + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (clarsimp simp: exposed_face_of_def) fix a b assume faceS: "S \ {x. a \ x = b} face_of S" and Ssub: "S \ {x. a \ x \ b}" show "\c d. S \ {x. c \ x \ d} \ S \ {x. a \ x = b} = S \ {x. c \ x = d} \ (S \ {x. a \ x = b} \ {} \ S \ {x. a \ x = b} \ S \ c \ 0) \ (S \ {x. a \ x = b} \ S \ (\w \ affine hull S. w + c \ affine hull S))" proof (cases "affine hull S \ {x. -a \ x \ -b} = {} \ affine hull S \ {x. - a \ x \ - b}") case True then show ?thesis proof assume "affine hull S \ {x. - a \ x \ - b} = {}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="1" in exI) using hull_subset by fastforce next assume "affine hull S \ {x. - a \ x \ - b}" then show ?thesis apply (rule_tac x="0" in exI) apply (rule_tac x="0" in exI) using Ssub hull_subset by fastforce qed next case False then obtain a' b' where "a' \ 0" and le: "affine hull S \ {x. a' \ x \ b'} = affine hull S \ {x. - a \ x \ - b}" and eq: "affine hull S \ {x. a' \ x = b'} = affine hull S \ {x. - a \ x = - b}" and mem: "\w. w \ affine hull S \ w + a' \ affine hull S" using affine_parallel_slice affine_affine_hull by metis show ?thesis proof (intro conjI impI allI ballI exI) have *: "S \ - (affine hull S \ {x. P x}) \ affine hull S \ {x. Q x} \ S \ {x. \ P x \ Q x}" for P Q using hull_subset by fastforce have "S \ {x. \ (a' \ x \ b') \ a' \ x = b'}" apply (rule *) apply (simp only: le eq) using Ssub by auto then show "S \ {x. - a' \ x \ - b'}" by auto show "S \ {x. a \ x = b} = S \ {x. - a' \ x = - b'}" using eq hull_subset [of S affine] by force show "\S \ {x. a \ x = b} \ {}; S \ {x. a \ x = b} \ S\ \ - a' \ 0" using \a' \ 0\ by auto show "w + - a' \ affine hull S" if "S \ {x. a \ x = b} \ S" "w \ affine hull S" for w proof - have "w + 1 *\<^sub>R (w - (w + a')) \ affine hull S" using affine_affine_hull mem mem_affine_3_minus that(2) by blast then show ?thesis by simp qed qed qed qed next assume ?rhs then show ?lhs unfolding exposed_face_of_def by blast qed subsection\Extreme points of a set: its singleton faces\ definition\<^marker>\tag important\ extreme_point_of :: "['a::real_vector, 'a set] \ bool" (infixr "(extreme'_point'_of)" 50) where "x extreme_point_of S \ x \ S \ (\a \ S. \b \ S. x \ open_segment a b)" lemma extreme_point_of_stillconvex: "convex S \ (x extreme_point_of S \ x \ S \ convex(S - {x}))" by (fastforce simp add: convex_contains_segment extreme_point_of_def open_segment_def) lemma face_of_singleton: "{x} face_of S \ x extreme_point_of S" by (fastforce simp add: extreme_point_of_def face_of_def) lemma extreme_point_not_in_REL_INTERIOR: fixes S :: "'a::real_normed_vector set" shows "\x extreme_point_of S; S \ {x}\ \ x \ rel_interior S" apply (simp add: face_of_singleton [symmetric]) apply (blast dest: face_of_disjoint_rel_interior) done lemma extreme_point_not_in_interior: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "x extreme_point_of S \ x \ interior S" apply (case_tac "S = {x}") apply (simp add: empty_interior_finite) by (meson contra_subsetD extreme_point_not_in_REL_INTERIOR interior_subset_rel_interior) lemma extreme_point_of_face: "F face_of S \ v extreme_point_of F \ v extreme_point_of S \ v \ F" by (meson empty_subsetI face_of_face face_of_singleton insert_subset) lemma extreme_point_of_convex_hull: "x extreme_point_of (convex hull S) \ x \ S" apply (simp add: extreme_point_of_stillconvex) using hull_minimal [of S "(convex hull S) - {x}" convex] using hull_subset [of S convex] apply blast done proposition extreme_points_of_convex_hull: "{x. x extreme_point_of (convex hull S)} \ S" using extreme_point_of_convex_hull by auto lemma extreme_point_of_empty [simp]: "\ (x extreme_point_of {})" by (simp add: extreme_point_of_def) lemma extreme_point_of_singleton [iff]: "x extreme_point_of {a} \ x = a" using extreme_point_of_stillconvex by auto lemma extreme_point_of_translation_eq: "(a + x) extreme_point_of (image (\x. a + x) S) \ x extreme_point_of S" by (auto simp: extreme_point_of_def) lemma extreme_points_of_translation: "{x. x extreme_point_of (image (\x. a + x) S)} = (\x. a + x) ` {x. x extreme_point_of S}" using extreme_point_of_translation_eq by auto (metis (no_types, lifting) image_iff mem_Collect_eq minus_add_cancel) lemma extreme_points_of_translation_subtract: "{x. x extreme_point_of (image (\x. x - a) S)} = (\x. x - a) ` {x. x extreme_point_of S}" using extreme_points_of_translation [of "- a" S] by simp lemma extreme_point_of_Int: "\x extreme_point_of S; x extreme_point_of T\ \ x extreme_point_of (S \ T)" by (simp add: extreme_point_of_def) lemma extreme_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" apply (simp add: face_of_singleton [symmetric]) by (metis face_of_Int_supporting_hyperplane_le_strong convex_singleton) lemma extreme_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ c extreme_point_of S" apply (simp add: face_of_singleton [symmetric]) by (metis face_of_Int_supporting_hyperplane_ge_strong convex_singleton) lemma exposed_point_of_Int_supporting_hyperplane_le: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" apply (simp add: exposed_face_of_def face_of_singleton) apply (force simp: extreme_point_of_Int_supporting_hyperplane_le) done lemma exposed_point_of_Int_supporting_hyperplane_ge: "\S \ {x. a \ x = b} = {c}; \x. x \ S \ a \ x \ b\ \ {c} exposed_face_of S" using exposed_point_of_Int_supporting_hyperplane_le [of S "-a" "-b" c] by simp lemma extreme_point_of_convex_hull_insert: "\finite S; a \ convex hull S\ \ a extreme_point_of (convex hull (insert a S))" apply (case_tac "a \ S") apply (simp add: hull_inc) using face_of_convex_hulls [of "insert a S" "{a}"] apply (auto simp: face_of_singleton hull_same) done subsection\Facets\ definition\<^marker>\tag important\ facet_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(facet'_of)" 50) where "F facet_of S \ F face_of S \ F \ {} \ aff_dim F = aff_dim S - 1" lemma facet_of_empty [simp]: "\ S facet_of {}" by (simp add: facet_of_def) lemma facet_of_irrefl [simp]: "\ S facet_of S " by (simp add: facet_of_def) lemma facet_of_imp_face_of: "F facet_of S \ F face_of S" by (simp add: facet_of_def) lemma facet_of_imp_subset: "F facet_of S \ F \ S" by (simp add: face_of_imp_subset facet_of_def) lemma hyperplane_facet_of_halfspace_le: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_ge hyperplane_face_of_halfspace_le DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_le) lemma hyperplane_facet_of_halfspace_ge: "a \ 0 \ {x. a \ x = b} facet_of {x. a \ x \ b}" unfolding facet_of_def hyperplane_eq_empty by (auto simp: hyperplane_face_of_halfspace_le hyperplane_face_of_halfspace_ge DIM_positive Suc_leI of_nat_diff aff_dim_halfspace_ge) lemma facet_of_halfspace_le: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" (is "?lhs = ?rhs") proof assume c: ?lhs with c facet_of_irrefl show ?rhs by (force simp: aff_dim_halfspace_le facet_of_def face_of_halfspace_le cong: conj_cong split: if_split_asm) next assume ?rhs then show ?lhs by (simp add: hyperplane_facet_of_halfspace_le) qed lemma facet_of_halfspace_ge: "F facet_of {x. a \ x \ b} \ a \ 0 \ F = {x. a \ x = b}" using facet_of_halfspace_le [of F "-a" "-b"] by simp subsection \Edges: faces of affine dimension 1\ (*FIXME too small subsection, rearrange? *) definition\<^marker>\tag important\ edge_of :: "['a::euclidean_space set, 'a set] \ bool" (infixr "(edge'_of)" 50) where "e edge_of S \ e face_of S \ aff_dim e = 1" lemma edge_of_imp_subset: "S edge_of T \ S \ T" by (simp add: edge_of_def face_of_imp_subset) subsection\Existence of extreme points\ proposition different_norm_3_collinear_points: fixes a :: "'a::euclidean_space" assumes "x \ open_segment a b" "norm(a) = norm(b)" "norm(x) = norm(b)" shows False proof - obtain u where "norm ((1 - u) *\<^sub>R a + u *\<^sub>R b) = norm b" and "a \ b" and u01: "0 < u" "u < 1" using assms by (auto simp: open_segment_image_interval if_splits) then have "(1 - u) *\<^sub>R a \ (1 - u) *\<^sub>R a + ((1 - u) * 2) *\<^sub>R a \ u *\<^sub>R b = (1 - u * u) *\<^sub>R (a \ a)" using assms by (simp add: norm_eq algebra_simps inner_commute) then have "(1 - u) *\<^sub>R ((1 - u) *\<^sub>R a \ a + (2 * u) *\<^sub>R a \ b) = (1 - u) *\<^sub>R ((1 + u) *\<^sub>R (a \ a))" by (simp add: algebra_simps) then have "(1 - u) *\<^sub>R (a \ a) + (2 * u) *\<^sub>R (a \ b) = (1 + u) *\<^sub>R (a \ a)" using u01 by auto then have "a \ b = a \ a" using u01 by (simp add: algebra_simps) then have "a = b" using \norm(a) = norm(b)\ norm_eq vector_eq by fastforce then show ?thesis using \a \ b\ by force qed proposition extreme_point_exists_convex: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "S \ {}" obtains x where "x extreme_point_of S" proof - obtain x where "x \ S" and xsup: "\y. y \ S \ norm y \ norm x" using distance_attains_sup [of S 0] assms by auto have False if "a \ S" "b \ S" and x: "x \ open_segment a b" for a b proof - have noax: "norm a \ norm x" and nobx: "norm b \ norm x" using xsup that by auto have "a \ b" using empty_iff open_segment_idem x by auto have *: "(1 - u) * na + u * nb < norm x" if "na < norm x" "nb \ norm x" "0 < u" "u < 1" for na nb u proof - have "(1 - u) * na + u * nb < (1 - u) * norm x + u * nb" by (simp add: that) also have "... \ (1 - u) * norm x + u * norm x" by (simp add: that) finally have "(1 - u) * na + u * nb < (1 - u) * norm x + u * norm x" . then show ?thesis using scaleR_collapse [symmetric, of "norm x" u] by auto qed have "norm x < norm x" if "norm a < norm x" using x apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) apply (rule norm_triangle_lt) apply (simp add: norm_mult) using * [of "norm a" "norm b"] nobx that apply blast done moreover have "norm x < norm x" if "norm b < norm x" using x apply (clarsimp simp only: open_segment_image_interval \a \ b\ if_False) apply (rule norm_triangle_lt) apply (simp add: norm_mult) using * [of "norm b" "norm a" "1-u" for u] noax that apply (simp add: add.commute) done ultimately have "\ (norm a < norm x) \ \ (norm b < norm x)" by auto then show ?thesis using different_norm_3_collinear_points noax nobx that(3) by fastforce qed then show ?thesis apply (rule_tac x=x in that) apply (force simp: extreme_point_of_def \x \ S\) done qed subsection\Krein-Milman, the weaker form\ proposition Krein_Milman: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = closure(convex hull {x. x extreme_point_of S})" proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" by (simp add: \compact S\ compact_imp_closed) have "closure (convex hull {x. x extreme_point_of S}) \ S" apply (rule closure_minimal [OF hull_minimal \closed S\]) using assms apply (auto simp: extreme_point_of_def) done moreover have "u \ closure (convex hull {x. x extreme_point_of S})" if "u \ S" for u proof (rule ccontr) assume unot: "u \ closure(convex hull {x. x extreme_point_of S})" then obtain a b where "a \ u < b" and ab: "\x. x \ closure(convex hull {x. x extreme_point_of S}) \ b < a \ x" using separating_hyperplane_closed_point [of "closure(convex hull {x. x extreme_point_of S})"] by blast have "continuous_on S ((\) a)" by (rule continuous_intros)+ then obtain m where "m \ S" and m: "\y. y \ S \ a \ m \ a \ y" using continuous_attains_inf [of S "\x. a \ x"] \compact S\ \u \ S\ by auto define T where "T = S \ {x. a \ x = a \ m}" have "m \ T" by (simp add: T_def \m \ S\) moreover have "compact T" by (simp add: T_def compact_Int_closed [OF \compact S\ closed_hyperplane]) moreover have "convex T" by (simp add: T_def convex_Int [OF \convex S\ convex_hyperplane]) ultimately obtain v where v: "v extreme_point_of T" using extreme_point_exists_convex [of T] by auto then have "{v} face_of T" by (simp add: face_of_singleton) also have "T face_of S" by (simp add: T_def m face_of_Int_supporting_hyperplane_ge [OF \convex S\]) finally have "v extreme_point_of S" by (simp add: face_of_singleton) then have "b < a \ v" using closure_subset by (simp add: closure_hull hull_inc ab) then show False using \a \ u < b\ \{v} face_of T\ face_of_imp_subset m T_def that by fastforce qed ultimately show ?thesis by blast qed text\Now the sharper form.\ lemma Krein_Milman_Minkowski_aux: fixes S :: "'a::euclidean_space set" assumes n: "dim S = n" and S: "compact S" "convex S" "0 \ S" shows "0 \ convex hull {x. x extreme_point_of S}" using n S proof (induction n arbitrary: S rule: less_induct) case (less n S) show ?case proof (cases "0 \ rel_interior S") case True with Krein_Milman show ?thesis by (metis subsetD convex_convex_hull convex_rel_interior_closure less.prems(2) less.prems(3) rel_interior_subset) next case False have "rel_interior S \ {}" by (simp add: rel_interior_convex_nonempty_aux less) then obtain c where c: "c \ rel_interior S" by blast obtain a where "a \ 0" and le_ay: "\y. y \ S \ a \ 0 \ a \ y" and less_ay: "\y. y \ rel_interior S \ a \ 0 < a \ y" by (blast intro: supporting_hyperplane_rel_boundary intro!: less False) have face: "S \ {x. a \ x = 0} face_of S" apply (rule face_of_Int_supporting_hyperplane_ge [OF \convex S\]) using le_ay by auto then have co: "compact (S \ {x. a \ x = 0})" "convex (S \ {x. a \ x = 0})" using less.prems by (blast intro: face_of_imp_compact face_of_imp_convex)+ have "a \ y = 0" if "y \ span (S \ {x. a \ x = 0})" for y proof - have "y \ span {x. a \ x = 0}" by (metis inf.cobounded2 span_mono subsetCE that) then show ?thesis by (blast intro: span_induct [OF _ subspace_hyperplane]) qed then have "dim (S \ {x. a \ x = 0}) < n" by (metis (no_types) less_ay c subsetD dim_eq_span inf.strict_order_iff inf_le1 \dim S = n\ not_le rel_interior_subset span_0 span_base) then have "0 \ convex hull {x. x extreme_point_of (S \ {x. a \ x = 0})}" by (rule less.IH) (auto simp: co less.prems) then show ?thesis by (metis (mono_tags, lifting) Collect_mono_iff \S \ {x. a \ x = 0} face_of S\ extreme_point_of_face hull_mono subset_iff) qed qed theorem Krein_Milman_Minkowski: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" shows "S = convex hull {x. x extreme_point_of S}" proof show "S \ convex hull {x. x extreme_point_of S}" proof fix a assume [simp]: "a \ S" have 1: "compact ((+) (- a) ` S)" by (simp add: \compact S\ compact_translation_subtract cong: image_cong_simp) have 2: "convex ((+) (- a) ` S)" by (simp add: \convex S\ compact_translation_subtract) show a_invex: "a \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski_aux [OF refl 1 2] convex_hull_translation [of "-a"] by (auto simp: extreme_points_of_translation_subtract translation_assoc cong: image_cong_simp) qed next show "convex hull {x. x extreme_point_of S} \ S" proof - have "{a. a extreme_point_of S} \ S" using extreme_point_of_def by blast then show ?thesis by (simp add: \convex S\ hull_minimal) qed qed subsection\Applying it to convex hulls of explicitly indicated finite sets\ corollary Krein_Milman_polytope: fixes S :: "'a::euclidean_space set" shows "finite S \ convex hull S = convex hull {x. x extreme_point_of (convex hull S)}" by (simp add: Krein_Milman_Minkowski finite_imp_compact_convex_hull) lemma extreme_points_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ {x. x extreme_point_of (convex hull S)} = S" by (metis (full_types) Krein_Milman_Minkowski compact_convex_hull convex_convex_hull extreme_points_of_convex_hull psubsetI) lemma extreme_point_of_convex_hull_eq: fixes S :: "'a::euclidean_space set" shows "\compact S; \T. T \ S \ convex hull T \ convex hull S\ \ (x extreme_point_of (convex hull S) \ x \ S)" using extreme_points_of_convex_hull_eq by auto lemma extreme_point_of_convex_hull_convex_independent: fixes S :: "'a::euclidean_space set" assumes "compact S" and S: "\a. a \ S \ a \ convex hull (S - {a})" shows "(x extreme_point_of (convex hull S) \ x \ S)" proof - have "convex hull T \ convex hull S" if "T \ S" for T proof - obtain a where "T \ S" "a \ S" "a \ T" using \T \ S\ by blast then show ?thesis by (metis (full_types) Diff_eq_empty_iff Diff_insert0 S hull_mono hull_subset insert_Diff_single subsetCE) qed then show ?thesis by (rule extreme_point_of_convex_hull_eq [OF \compact S\]) qed lemma extreme_point_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" shows "\ affine_dependent S \ (x extreme_point_of (convex hull S) \ x \ S)" by (metis aff_independent_finite affine_dependent_def affine_hull_convex_hull extreme_point_of_convex_hull_convex_independent finite_imp_compact hull_inc) text\Elementary proofs exist, not requiring Euclidean spaces and all this development\ lemma extreme_point_of_convex_hull_2: fixes x :: "'a::euclidean_space" shows "x extreme_point_of (convex hull {a,b}) \ x = a \ x = b" proof - have "x extreme_point_of (convex hull {a,b}) \ x \ {a,b}" by (intro extreme_point_of_convex_hull_affine_independent affine_independent_2) then show ?thesis by simp qed lemma extreme_point_of_segment: fixes x :: "'a::euclidean_space" shows "x extreme_point_of closed_segment a b \ x = a \ x = b" by (simp add: extreme_point_of_convex_hull_2 segment_convex_hull) lemma face_of_convex_hull_subset: fixes S :: "'a::euclidean_space set" assumes "compact S" and T: "T face_of (convex hull S)" obtains s' where "s' \ S" "T = convex hull s'" apply (rule_tac s' = "{x. x extreme_point_of T}" in that) using T extreme_point_of_convex_hull extreme_point_of_face apply blast by (metis (no_types) Krein_Milman_Minkowski assms compact_convex_hull convex_convex_hull face_of_imp_compact face_of_imp_convex) lemma face_of_convex_hull_aux: assumes eq: "x *\<^sub>R p = u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c" and x: "u + v + w = x" "x \ 0" and S: "affine S" "a \ S" "b \ S" "c \ S" shows "p \ S" proof - have "p = (u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x" by (metis \x \ 0\ eq mult.commute right_inverse scaleR_one scaleR_scaleR) moreover have "affine hull {a,b,c} \ S" by (simp add: S hull_minimal) moreover have "(u *\<^sub>R a + v *\<^sub>R b + w *\<^sub>R c) /\<^sub>R x \ affine hull {a,b,c}" apply (simp add: affine_hull_3) apply (rule_tac x="u/x" in exI) apply (rule_tac x="v/x" in exI) apply (rule_tac x="w/x" in exI) using x apply (auto simp: algebra_simps field_split_simps) done ultimately show ?thesis by force qed proposition face_of_convex_hull_insert_eq: fixes a :: "'a :: euclidean_space" assumes "finite S" and a: "a \ affine hull S" shows "(F face_of (convex hull (insert a S)) \ F face_of (convex hull S) \ (\F'. F' face_of (convex hull S) \ F = convex hull (insert a F')))" (is "F face_of ?CAS \ _") proof safe assume F: "F face_of ?CAS" and *: "\F'. F' face_of convex hull S \ F = convex hull insert a F'" obtain T where T: "T \ insert a S" and FeqT: "F = convex hull T" by (metis F \finite S\ compact_insert finite_imp_compact face_of_convex_hull_subset) show "F face_of convex hull S" proof (cases "a \ T") case True have "F = convex hull insert a (convex hull T \ convex hull S)" proof have "T \ insert a (convex hull T \ convex hull S)" using T hull_subset by fastforce then show "F \ convex hull insert a (convex hull T \ convex hull S)" by (simp add: FeqT hull_mono) show "convex hull insert a (convex hull T \ convex hull S) \ F" apply (rule hull_minimal) using True by (auto simp: \F = convex hull T\ hull_inc) qed moreover have "convex hull T \ convex hull S face_of convex hull S" by (metis F FeqT convex_convex_hull face_of_slice hull_mono inf.absorb_iff2 subset_insertI) ultimately show ?thesis using * by force next case False then show ?thesis by (metis FeqT F T face_of_subset hull_mono subset_insert subset_insertI) qed next assume "F face_of convex hull S" show "F face_of ?CAS" by (simp add: \F face_of convex hull S\ a face_of_convex_hull_insert \finite S\) next fix F assume F: "F face_of convex hull S" show "convex hull insert a F face_of ?CAS" proof (cases "S = {}") case True then show ?thesis using F face_of_affine_eq by auto next case False have anotc: "a \ convex hull S" by (metis (no_types) a affine_hull_convex_hull hull_inc) show ?thesis proof (cases "F = {}") case True show ?thesis using anotc by (simp add: \F = {}\ \finite S\ extreme_point_of_convex_hull_insert face_of_singleton) next case False have "convex hull insert a F \ ?CAS" by (simp add: F a \finite S\ convex_hull_subset face_of_convex_hull_insert face_of_imp_subset hull_inc) moreover have "(\y v. (1 - ub) *\<^sub>R a + ub *\<^sub>R b = (1 - v) *\<^sub>R a + v *\<^sub>R y \ 0 \ v \ v \ 1 \ y \ F) \ (\x u. (1 - uc) *\<^sub>R a + uc *\<^sub>R c = (1 - u) *\<^sub>R a + u *\<^sub>R x \ 0 \ u \ u \ 1 \ x \ F)" if *: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x \ open_segment ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 \ ub" "ub \ 1" "0 \ uc" "uc \ 1" "0 \ ux" "ux \ 1" and b: "b \ convex hull S" and c: "c \ convex hull S" and "x \ F" for b c ub uc ux x proof - obtain v where ne: "(1 - ub) *\<^sub>R a + ub *\<^sub>R b \ (1 - uc) *\<^sub>R a + uc *\<^sub>R c" and eq: "(1 - ux) *\<^sub>R a + ux *\<^sub>R x = (1 - v) *\<^sub>R ((1 - ub) *\<^sub>R a + ub *\<^sub>R b) + v *\<^sub>R ((1 - uc) *\<^sub>R a + uc *\<^sub>R c)" and "0 < v" "v < 1" using * by (auto simp: in_segment) then have 0: "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a + (ux *\<^sub>R x - (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)) = 0" by (auto simp: algebra_simps) then have "((1 - ux) - ((1 - v) * (1 - ub) + v * (1 - uc))) *\<^sub>R a = ((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c + (-ux) *\<^sub>R x" by (auto simp: algebra_simps) then have "a \ affine hull S" if "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) \ 0" apply (rule face_of_convex_hull_aux) using b c that apply (auto simp: algebra_simps) using F convex_hull_subset_affine_hull face_of_imp_subset \x \ F\ apply blast+ done then have "1 - ux - ((1 - v) * (1 - ub) + v * (1 - uc)) = 0" using a by blast with 0 have equx: "(1 - v) * ub + v * uc = ux" and uxx: "ux *\<^sub>R x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c)" by auto (auto simp: algebra_simps) show ?thesis proof (cases "uc = 0") case True then show ?thesis using equx 0 \0 \ ub\ \ub \ 1\ \v < 1\ \x \ F\ apply (auto simp: algebra_simps) apply (rule_tac x=x in exI, simp) apply (rule_tac x=ub in exI, auto) apply (metis add.left_neutral diff_eq_eq less_irrefl mult.commute mult_cancel_right1 real_vector.scale_cancel_left real_vector.scale_left_diff_distrib) using \x \ F\ \uc \ 1\ apply blast done next case False show ?thesis proof (cases "ub = 0") case True then show ?thesis using equx 0 \0 \ uc\ \uc \ 1\ \0 < v\ \x \ F\ \uc \ 0\ by (force simp: algebra_simps) next case False then have "0 < ub" "0 < uc" using \uc \ 0\ \0 \ ub\ \0 \ uc\ by auto then have "ux \ 0" by (metis \0 < v\ \v < 1\ diff_ge_0_iff_ge dual_order.strict_implies_order equx leD le_add_same_cancel2 zero_le_mult_iff zero_less_mult_iff) have "b \ F \ c \ F" proof (cases "b = c") case True then show ?thesis by (metis \ux \ 0\ equx real_vector.scale_cancel_left scaleR_add_left uxx \x \ F\) next case False have "x = (((1 - v) * ub) *\<^sub>R b + (v * uc) *\<^sub>R c) /\<^sub>R ux" by (metis \ux \ 0\ uxx mult.commute right_inverse scaleR_one scaleR_scaleR) also have "... = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" using \ux \ 0\ equx apply (auto simp: algebra_simps field_split_simps) by (metis add.commute add_diff_eq add_divide_distrib diff_add_cancel scaleR_add_left) finally have "x = (1 - v * uc / ux) *\<^sub>R b + (v * uc / ux) *\<^sub>R c" . then have "x \ open_segment b c" apply (simp add: in_segment \b \ c\) apply (rule_tac x="(v * uc) / ux" in exI) using \0 \ ux\ \ux \ 0\ \0 < uc\ \0 < v\ \0 < ub\ \v < 1\ equx apply (force simp: algebra_simps field_split_simps) done then show ?thesis by (rule face_ofD [OF F _ b c \x \ F\]) qed with \0 \ ub\ \ub \ 1\ \0 \ uc\ \uc \ 1\ show ?thesis by blast qed qed qed moreover have "convex hull F = F" by (meson F convex_hull_eq face_of_imp_convex) ultimately show ?thesis unfolding face_of_def by (fastforce simp: convex_hull_insert_alt \S \ {}\ \F \ {}\) qed qed qed lemma face_of_convex_hull_insert2: fixes a :: "'a :: euclidean_space" assumes S: "finite S" and a: "a \ affine hull S" and F: "F face_of convex hull S" shows "convex hull (insert a F) face_of convex hull (insert a S)" by (metis F face_of_convex_hull_insert_eq [OF S a]) proposition face_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(T face_of (convex hull S) \ (\c. c \ S \ T = convex hull c))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson \T face_of convex hull S\ aff_independent_finite assms face_of_convex_hull_subset finite_imp_compact) next assume ?rhs then obtain c where "c \ S" and T: "T = convex hull c" by blast have "affine hull c \ affine hull (S - c) = {}" apply (rule disjoint_affine_hull [OF assms \c \ S\], auto) done then have "affine hull c \ convex hull (S - c) = {}" using convex_hull_subset_affine_hull by fastforce then show ?lhs by (metis face_of_convex_hulls \c \ S\ aff_independent_finite assms T) qed lemma facet_of_convex_hull_affine_independent: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "T facet_of (convex hull S) \ T \ {} \ (\u. u \ S \ T = convex hull (S - {u}))" (is "?lhs = ?rhs") proof assume ?lhs then have "T face_of (convex hull S)" "T \ {}" and afft: "aff_dim T = aff_dim (convex hull S) - 1" by (auto simp: facet_of_def) then obtain c where "c \ S" and c: "T = convex hull c" by (auto simp: face_of_convex_hull_affine_independent [OF assms]) then have affs: "aff_dim S = aff_dim c + 1" by (metis aff_dim_convex_hull afft eq_diff_eq) have "\ affine_dependent c" using \c \ S\ affine_dependent_subset assms by blast with affs have "card (S - c) = 1" apply (simp add: aff_dim_affine_independent [symmetric] aff_dim_convex_hull) by (metis aff_dim_affine_independent aff_independent_finite One_nat_def \c \ S\ add.commute add_diff_cancel_right' assms card_Diff_subset card_mono of_nat_1 of_nat_diff of_nat_eq_iff) then obtain u where u: "u \ S - c" by (metis DiffI \c \ S\ aff_independent_finite assms cancel_comm_monoid_add_class.diff_cancel card_Diff_subset subsetI subset_antisym zero_neq_one) then have u: "S = insert u c" by (metis Diff_subset \c \ S\ \card (S - c) = 1\ card_1_singletonE double_diff insert_Diff insert_subset singletonD) have "T = convex hull (c - {u})" by (metis Diff_empty Diff_insert0 \T facet_of convex hull S\ c facet_of_irrefl insert_absorb u) with \T \ {}\ show ?rhs using c u by auto next assume ?rhs then obtain u where "T \ {}" "u \ S" and u: "T = convex hull (S - {u})" by (force simp: facet_of_def) then have "\ S \ {u}" using \T \ {}\ u by auto have [simp]: "aff_dim (convex hull (S - {u})) = aff_dim (convex hull S) - 1" using assms \u \ S\ apply (simp add: aff_dim_convex_hull affine_dependent_def) apply (drule bspec, assumption) by (metis add_diff_cancel_right' aff_dim_insert insert_Diff [of u S]) show ?lhs apply (subst u) apply (simp add: \\ S \ {u}\ facet_of_def face_of_convex_hull_affine_independent [OF assms], blast) done qed lemma facet_of_convex_hull_affine_independent_alt: fixes S :: "'a::euclidean_space set" shows "\affine_dependent S \ (T facet_of (convex hull S) \ 2 \ card S \ (\u. u \ S \ T = convex hull (S - {u})))" apply (simp add: facet_of_convex_hull_affine_independent) apply (auto simp: Set.subset_singleton_iff) apply (metis Diff_cancel Int_empty_right Int_insert_right_if1 aff_independent_finite card_eq_0_iff card_insert_if card_mono card_subset_eq convex_hull_eq_empty eq_iff equals0D finite_insert finite_subset inf.absorb_iff2 insert_absorb insert_not_empty not_less_eq_eq numeral_2_eq_2) done lemma segment_face_of: assumes "(closed_segment a b) face_of S" shows "a extreme_point_of S" "b extreme_point_of S" proof - have as: "{a} face_of S" by (metis (no_types) assms convex_hull_singleton empty_iff extreme_point_of_convex_hull_insert face_of_face face_of_singleton finite.emptyI finite.insertI insert_absorb insert_iff segment_convex_hull) moreover have "{b} face_of S" proof - have "b \ convex hull {a} \ b extreme_point_of convex hull {b, a}" by (meson extreme_point_of_convex_hull_insert finite.emptyI finite.insertI) moreover have "closed_segment a b = convex hull {b, a}" using closed_segment_commute segment_convex_hull by blast ultimately show ?thesis by (metis as assms face_of_face convex_hull_singleton empty_iff face_of_singleton insertE) qed ultimately show "a extreme_point_of S" "b extreme_point_of S" using face_of_singleton by blast+ qed proposition Krein_Milman_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "compact S" shows "S = convex hull (frontier S)" (is "?lhs = ?rhs") proof have "?lhs \ convex hull {x. x extreme_point_of S}" using Krein_Milman_Minkowski assms by blast also have "... \ ?rhs" apply (rule hull_mono) apply (auto simp: frontier_def extreme_point_not_in_interior) using closure_subset apply (force simp: extreme_point_of_def) done finally show "?lhs \ ?rhs" . next have "?rhs \ convex hull S" by (metis Diff_subset \compact S\ closure_closed compact_eq_bounded_closed frontier_def hull_mono) also have "... \ ?lhs" by (simp add: \convex S\ hull_same) finally show "?rhs \ ?lhs" . qed subsection\Polytopes\ definition\<^marker>\tag important\ polytope where "polytope S \ \v. finite v \ S = convex hull v" lemma polytope_translation_eq: "polytope (image (\x. a + x) S) \ polytope S" apply (simp add: polytope_def, safe) apply (metis convex_hull_translation finite_imageI translation_galois) by (metis convex_hull_translation finite_imageI) lemma polytope_linear_image: "\linear f; polytope p\ \ polytope(image f p)" unfolding polytope_def using convex_hull_linear_image by blast lemma polytope_empty: "polytope {}" using convex_hull_empty polytope_def by blast lemma polytope_convex_hull: "finite S \ polytope(convex hull S)" using polytope_def by auto lemma polytope_Times: "\polytope S; polytope T\ \ polytope(S \ T)" unfolding polytope_def by (metis finite_cartesian_product convex_hull_Times) lemma face_of_polytope_polytope: fixes S :: "'a::euclidean_space set" shows "\polytope S; F face_of S\ \ polytope F" unfolding polytope_def by (meson face_of_convex_hull_subset finite_imp_compact finite_subset) lemma finite_polytope_faces: fixes S :: "'a::euclidean_space set" assumes "polytope S" shows "finite {F. F face_of S}" proof - obtain v where "finite v" "S = convex hull v" using assms polytope_def by auto have "finite ((hull) convex ` {T. T \ v})" by (simp add: \finite v\) moreover have "{F. F face_of S} \ ((hull) convex ` {T. T \ v})" by (metis (no_types, lifting) \finite v\ \S = convex hull v\ face_of_convex_hull_subset finite_imp_compact image_eqI mem_Collect_eq subsetI) ultimately show ?thesis by (blast intro: finite_subset) qed lemma finite_polytope_facets: assumes "polytope S" shows "finite {T. T facet_of S}" by (simp add: assms facet_of_def finite_polytope_faces) lemma polytope_scaling: assumes "polytope S" shows "polytope (image (\x. c *\<^sub>R x) S)" by (simp add: assms polytope_linear_image) lemma polytope_imp_compact: fixes S :: "'a::real_normed_vector set" shows "polytope S \ compact S" by (metis finite_imp_compact_convex_hull polytope_def) lemma polytope_imp_convex: "polytope S \ convex S" by (metis convex_convex_hull polytope_def) lemma polytope_imp_closed: fixes S :: "'a::real_normed_vector set" shows "polytope S \ closed S" by (simp add: compact_imp_closed polytope_imp_compact) lemma polytope_imp_bounded: fixes S :: "'a::real_normed_vector set" shows "polytope S \ bounded S" by (simp add: compact_imp_bounded polytope_imp_compact) lemma polytope_interval: "polytope(cbox a b)" unfolding polytope_def by (meson closed_interval_as_convex_hull) lemma polytope_sing: "polytope {a}" using polytope_def by force lemma face_of_polytope_insert: "\polytope S; a \ affine hull S; F face_of S\ \ F face_of convex hull (insert a S)" by (metis (no_types, lifting) affine_hull_convex_hull face_of_convex_hull_insert hull_insert polytope_def) proposition face_of_polytope_insert2: fixes a :: "'a :: euclidean_space" assumes "polytope S" "a \ affine hull S" "F face_of S" shows "convex hull (insert a F) face_of convex hull (insert a S)" proof - obtain V where "finite V" "S = convex hull V" using assms by (auto simp: polytope_def) then have "convex hull (insert a F) face_of convex hull (insert a V)" using affine_hull_convex_hull assms face_of_convex_hull_insert2 by blast then show ?thesis by (metis \S = convex hull V\ hull_insert) qed subsection\Polyhedra\ definition\<^marker>\tag important\ polyhedron where "polyhedron S \ \F. finite F \ S = \ F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b})" lemma polyhedron_Int [intro,simp]: "\polyhedron S; polyhedron T\ \ polyhedron (S \ T)" apply (simp add: polyhedron_def, clarify) apply (rename_tac F G) apply (rule_tac x="F \ G" in exI, auto) done lemma polyhedron_UNIV [iff]: "polyhedron UNIV" unfolding polyhedron_def by (rule_tac x="{}" in exI) auto lemma polyhedron_Inter [intro,simp]: "\finite F; \S. S \ F \ polyhedron S\ \ polyhedron(\F)" by (induction F rule: finite_induct) auto lemma polyhedron_empty [iff]: "polyhedron ({} :: 'a :: euclidean_space set)" proof - have "\a. a \ 0 \ (\b. {x. (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b})" by (rule_tac x="(SOME i. i \ Basis)" in exI) (force simp: SOME_Basis nonzero_Basis) moreover have "\a b. a \ 0 \ {x. - (SOME i. i \ Basis) \ x \ - 1} = {x. a \ x \ b}" apply (rule_tac x="-(SOME i. i \ Basis)" in exI) apply (rule_tac x="-1" in exI) apply (simp add: SOME_Basis nonzero_Basis) done ultimately show ?thesis unfolding polyhedron_def apply (rule_tac x="{{x. (SOME i. i \ Basis) \ x \ -1}, {x. -(SOME i. i \ Basis) \ x \ -1}}" in exI) apply force done qed lemma polyhedron_halfspace_le: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" proof (cases "a = 0") case True then show ?thesis by auto next case False then show ?thesis unfolding polyhedron_def by (rule_tac x="{{x. a \ x \ b}}" in exI) auto qed lemma polyhedron_halfspace_ge: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x \ b}" using polyhedron_halfspace_le [of "-a" "-b"] by simp lemma polyhedron_hyperplane: fixes a :: "'a :: euclidean_space" shows "polyhedron {x. a \ x = b}" proof - have "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by force then show ?thesis by (simp add: polyhedron_halfspace_ge polyhedron_halfspace_le) qed lemma affine_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "affine S \ polyhedron S" by (metis affine_hull_eq polyhedron_Inter polyhedron_hyperplane affine_hull_finite_intersection_hyperplanes [of S]) lemma polyhedron_imp_closed: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S" apply (simp add: polyhedron_def) using closed_halfspace_le by fastforce lemma polyhedron_imp_convex: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ convex S" apply (simp add: polyhedron_def) using convex_Inter convex_halfspace_le by fastforce lemma polyhedron_affine_hull: fixes S :: "'a :: euclidean_space set" shows "polyhedron(affine hull S)" by (simp add: affine_imp_polyhedron) subsection\Canonical polyhedron representation making facial structure explicit\ proposition polyhedron_Int_affine: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: polyhedron_def) apply (erule ex_forward) using hull_subset apply force done next assume ?rhs then show ?lhs apply clarify apply (erule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le) done qed proposition rel_interior_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "rel_interior S = {x \ S. \h \ F. a h \ x < b h}" proof - have rels: "\x. x \ rel_interior S \ x \ S" by (meson IntE mem_rel_interior) moreover have "a i \ x < b i" if x: "x \ rel_interior S" and "i \ F" for x i proof - have fif: "F - {i} \ F" using \i \ F\ Diff_insert_absorb Diff_subset set_insert psubsetI by blast then have "S \ affine hull S \ \(F - {i})" by (rule psub) then obtain z where ssub: "S \ \(F - {i})" and zint: "z \ \(F - {i})" and "z \ S" and zaff: "z \ affine hull S" by auto have "z \ x" using \z \ S\ rels x by blast have "z \ affine hull S \ \F" using \z \ S\ seq by auto then have aiz: "a i \ z > b i" using faceq zint zaff by fastforce obtain e where "e > 0" "x \ S" and e: "ball x e \ affine hull S \ S" using x by (auto simp: mem_rel_interior_ball) then have ins: "\y. \norm (x - y) < e; y \ affine hull S\ \ y \ S" by (metis IntI subsetD dist_norm mem_ball) define \ where "\ = min (1/2) (e / 2 / norm(z - x))" have "norm (\ *\<^sub>R x - \ *\<^sub>R z) = norm (\ *\<^sub>R (x - z))" by (simp add: \_def algebra_simps norm_mult) also have "... = \ * norm (x - z)" using \e > 0\ by (simp add: \_def) also have "... < e" using \z \ x\ \e > 0\ by (simp add: \_def min_def field_split_simps norm_minus_commute) finally have lte: "norm (\ *\<^sub>R x - \ *\<^sub>R z) < e" . have \_aff: "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ affine hull S" by (metis \x \ S\ add.commute affine_affine_hull diff_add_cancel hull_inc mem_affine zaff) have "\ *\<^sub>R z + (1 - \) *\<^sub>R x \ S" apply (rule ins [OF _ \_aff]) apply (simp add: algebra_simps lte) done then obtain l where l: "0 < l" "l < 1" and ls: "(l *\<^sub>R z + (1 - l) *\<^sub>R x) \ S" apply (rule_tac l = \ in that) using \e > 0\ \z \ x\ apply (auto simp: \_def) done then have i: "l *\<^sub>R z + (1 - l) *\<^sub>R x \ i" using seq \i \ F\ by auto have "b i * l + (a i \ x) * (1 - l) < a i \ (l *\<^sub>R z + (1 - l) *\<^sub>R x)" using l by (simp add: algebra_simps aiz) also have "\ \ b i" using i l using faceq mem_Collect_eq \i \ F\ by blast finally have "(a i \ x) * (1 - l) < b i * (1 - l)" by (simp add: algebra_simps) with l show ?thesis by simp qed moreover have "x \ rel_interior S" if "x \ S" and less: "\h. h \ F \ a h \ x < b h" for x proof - have 1: "\h. h \ F \ x \ interior h" by (metis interior_halfspace_le mem_Collect_eq less faceq) have 2: "\y. \\h\F. y \ interior h; y \ affine hull S\ \ y \ S" by (metis IntI Inter_iff contra_subsetD interior_subset seq) show ?thesis apply (simp add: rel_interior \x \ S\) apply (rule_tac x="\h\F. interior h" in exI) apply (auto simp: \finite F\ open_INT 1 2) done qed ultimately show ?thesis by blast qed lemma polyhedron_Int_affine_parallel: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain F where "finite F" and seq: "S = (affine hull S) \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" by (fastforce simp add: polyhedron_Int_affine) then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?rhs proof - have "\a' b'. a' \ 0 \ affine hull S \ {x. a' \ x \ b'} = affine hull S \ h \ (\w \ affine hull S. (w + a') \ affine hull S)" if "h \ F" "\(affine hull S \ h)" for h proof - have "a h \ 0" and "h = {x. a h \ x \ b h}" "h \ \F = \F" using \h \ F\ ab by auto then have "(affine hull S) \ {x. a h \ x \ b h} \ {}" by (metis (no_types) affine_hull_eq_empty inf.absorb_iff2 inf_assoc inf_bot_right inf_commute seq that(2)) moreover have "\ (affine hull S \ {x. a h \ x \ b h})" using \h = {x. a h \ x \ b h}\ that(2) by blast ultimately show ?thesis using affine_parallel_slice [of "affine hull S"] by (metis \h = {x. a h \ x \ b h}\ affine_affine_hull) qed then obtain a b where ab: "\h. \h \ F; \ (affine hull S \ h)\ \ a h \ 0 \ affine hull S \ {x. a h \ x \ b h} = affine hull S \ h \ (\w \ affine hull S. (w + a h) \ affine hull S)" by metis have seq2: "S = affine hull S \ (\h\{h \ F. \ affine hull S \ h}. {x. a h \ x \ b h})" by (subst seq) (auto simp: ab INT_extend_simps) show ?thesis apply (rule_tac x="(\h. {x. a h \ x \ b h}) ` {h. h \ F \ \(affine hull S \ h)}" in exI) apply (intro conjI seq2) using \finite F\ apply force using ab apply blast done qed next assume ?rhs then show ?lhs apply (simp add: polyhedron_Int_affine) by metis qed proposition polyhedron_Int_affine_parallel_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ (\F) \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)) \ (\F'. F' \ F \ S \ (affine hull S) \ (\F')))" (is "?lhs = ?rhs") proof assume ?lhs then obtain f0 where f0: "finite f0" "S = (affine hull S) \ (\f0)" (is "?P f0") "\h \ f0. \a b. a \ 0 \ h = {x. a \ x \ b} \ (\x \ affine hull S. (x + a) \ affine hull S)" (is "?Q f0") by (force simp: polyhedron_Int_affine_parallel) define n where "n = (LEAST n. \F. card F = n \ finite F \ ?P F \ ?Q F)" have nf: "\F. card F = n \ finite F \ ?P F \ ?Q F" apply (simp add: n_def) apply (rule LeastI [where k = "card f0"]) using f0 apply auto done then obtain F where F: "card F = n" "finite F" and seq: "?P F" and aff: "?Q F" by blast then have "\ (finite g \ ?P g \ ?Q g)" if "card g < n" for g using that by (auto simp: n_def dest!: not_less_Least) then have *: "\ (?P g \ ?Q g)" if "g \ F" for g using that \finite F\ psubset_card_mono \card F = n\ by (metis finite_Int inf.strict_order_iff) have 1: "\F'. F' \ F \ S \ affine hull S \ \F'" by (subst seq) blast have 2: "\F'. F' \ F \ S \ affine hull S \ \F'" apply (frule *) by (metis aff subsetCE subset_iff_psubset_eq) show ?rhs by (metis \finite F\ seq aff psubsetI 1 2) next assume ?rhs then show ?lhs by (auto simp: polyhedron_Int_affine_parallel) qed lemma polyhedron_Int_affine_minimal: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ (\F. finite F \ S = (affine hull S) \ \F \ (\h \ F. \a b. a \ 0 \ h = {x. a \ x \ b}) \ (\F'. F' \ F \ S \ (affine hull S) \ \F'))" apply (rule iffI) apply (force simp: polyhedron_Int_affine_parallel_minimal elim!: ex_forward) apply (auto simp: polyhedron_Int_affine elim!: ex_forward) done proposition facet_of_polyhedron_explicit: assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" shows "c facet_of S \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" proof (cases "S = {}") case True with psub show ?thesis by force next case False have "polyhedron S" apply (simp add: polyhedron_Int_affine) apply (rule_tac x=F in exI) using assms apply force done then have "convex S" by (rule polyhedron_imp_convex) with False rel_interior_eq_empty have "rel_interior S \ {}" by blast then obtain x where "x \ rel_interior S" by auto then obtain T where "open T" "x \ T" "x \ S" "T \ affine hull S \ S" by (force simp: mem_rel_interior) then have xaff: "x \ affine hull S" and xint: "x \ \F" using seq hull_inc by auto have "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) with \x \ rel_interior S\ have [simp]: "\h. h\F \ a h \ x < b h" by blast have *: "(S \ {x. a h \ x = b h}) facet_of S" if "h \ F" for h proof - have "S \ affine hull S \ \(F - {h})" using psub that by (metis Diff_disjoint Diff_subset insert_disjoint(2) psubsetI) then obtain z where zaff: "z \ affine hull S" and zint: "z \ \(F - {h})" and "z \ S" by force then have "z \ x" "z \ h" using seq \x \ S\ by auto have "x \ h" using that xint by auto then have able: "a h \ x \ b h" using faceq that by blast also have "... < a h \ z" using \z \ h\ faceq [OF that] xint by auto finally have xltz: "a h \ x < a h \ z" . define l where "l = (b h - a h \ x) / (a h \ z - a h \ x)" define w where "w = (1 - l) *\<^sub>R x + l *\<^sub>R z" have "0 < l" "l < 1" using able xltz \b h < a h \ z\ \h \ F\ by (auto simp: l_def field_split_simps) have awlt: "a i \ w < b i" if "i \ F" "i \ h" for i proof - have "(1 - l) * (a i \ x) < (1 - l) * b i" by (simp add: \l < 1\ \i \ F\) moreover have "l * (a i \ z) \ l * b i" apply (rule mult_left_mono) apply (metis Diff_insert_absorb Inter_iff Set.set_insert \h \ F\ faceq insertE mem_Collect_eq that zint) using \0 < l\ apply simp done ultimately show ?thesis by (simp add: w_def algebra_simps) qed have weq: "a h \ w = b h" using xltz unfolding w_def l_def by (simp add: algebra_simps) (simp add: field_simps) have "w \ affine hull S" by (simp add: w_def mem_affine xaff zaff) moreover have "w \ \F" using \a h \ w = b h\ awlt faceq less_eq_real_def by blast ultimately have "w \ S" using seq by blast with weq have "S \ {x. a h \ x = b h} \ {}" by blast moreover have "S \ {x. a h \ x = b h} face_of S" apply (rule face_of_Int_supporting_hyperplane_le) apply (rule \convex S\) apply (subst (asm) seq) using faceq that apply fastforce done moreover have "affine hull (S \ {x. a h \ x = b h}) = (affine hull S) \ {x. a h \ x = b h}" proof show "affine hull (S \ {x. a h \ x = b h}) \ affine hull S \ {x. a h \ x = b h}" apply (intro Int_greatest hull_mono Int_lower1) apply (metis affine_hull_eq affine_hyperplane hull_mono inf_le2) done next show "affine hull S \ {x. a h \ x = b h} \ affine hull (S \ {x. a h \ x = b h})" proof fix y assume yaff: "y \ affine hull S \ {y. a h \ y = b h}" obtain T where "0 < T" and T: "\j. \j \ F; j \ h\ \ T * (a j \ y - a j \ w) \ b j - a j \ w" proof (cases "F - {h} = {}") case True then show ?thesis by (rule_tac T=1 in that) auto next case False then obtain h' where h': "h' \ F - {h}" by auto let ?body = "(\j. if 0 < a j \ y - a j \ w then (b j - a j \ w) / (a j \ y - a j \ w) else 1) ` (F - {h})" define inff where "inff = Inf ?body" from \finite F\ have "finite ?body" by blast moreover from h' have "?body \ {}" by blast moreover have "j > 0" if "j \ ?body" for j proof - from that obtain x where "x \ F" and "x \ h" and *: "j = (if 0 < a x \ y - a x \ w then (b x - a x \ w) / (a x \ y - a x \ w) else 1)" by blast with awlt [of x] have "a x \ w < b x" by simp with * show ?thesis by simp qed ultimately have "0 < inff" by (simp_all add: finite_less_Inf_iff inff_def) moreover have "inff * (a j \ y - a j \ w) \ b j - a j \ w" if "j \ F" "j \ h" for j proof (cases "a j \ w < a j \ y") case True then have "inff \ (b j - a j \ w) / (a j \ y - a j \ w)" apply (simp add: inff_def) apply (rule cInf_le_finite) using \finite F\ apply blast apply (simp add: that split: if_split_asm) done then show ?thesis using \0 < inff\ awlt [OF that] mult_strict_left_mono by (fastforce simp add: algebra_simps field_split_simps split: if_split_asm) next case False with \0 < inff\ have "inff * (a j \ y - a j \ w) \ 0" by (simp add: mult_le_0_iff) also have "... < b j - a j \ w" by (simp add: awlt that) finally show ?thesis by simp qed ultimately show ?thesis by (blast intro: that) qed define c where "c = (1 - T) *\<^sub>R w + T *\<^sub>R y" have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ j" if "j \ F" for j proof (cases "j = h") case True have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a h \ x \ b h}" using weq yaff by (auto simp: algebra_simps) with True faceq [OF that] show ?thesis by metis next case False with T that have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ {x. a j \ x \ b j}" by (simp add: algebra_simps) with faceq [OF that] show ?thesis by simp qed moreover have "(1 - T) *\<^sub>R w + T *\<^sub>R y \ affine hull S" apply (rule affine_affine_hull [simplified affine_alt, rule_format]) apply (simp add: \w \ affine hull S\) using yaff apply blast done ultimately have "c \ S" using seq by (force simp: c_def) moreover have "a h \ c = b h" using yaff by (force simp: c_def algebra_simps weq) ultimately have caff: "c \ affine hull (S \ {y. a h \ y = b h})" by (simp add: hull_inc) have waff: "w \ affine hull (S \ {y. a h \ y = b h})" using \w \ S\ weq by (blast intro: hull_inc) have yeq: "y = (1 - inverse T) *\<^sub>R w + c /\<^sub>R T" using \0 < T\ by (simp add: c_def algebra_simps) show "y \ affine hull (S \ {y. a h \ y = b h})" by (metis yeq affine_affine_hull [simplified affine_alt, rule_format, OF waff caff]) qed qed ultimately show ?thesis apply (simp add: facet_of_def) apply (subst aff_dim_affine_hull [symmetric]) using \b h < a h \ z\ zaff apply (force simp: aff_dim_affine_Int_hyperplane) done qed show ?thesis proof show "\h. h \ F \ c = S \ {x. a h \ x = b h} \ c facet_of S" using * by blast next assume "c facet_of S" then have "c face_of S" "convex c" "c \ {}" and affc: "aff_dim c = aff_dim S - 1" by (auto simp: facet_of_def face_of_imp_convex) then obtain x where x: "x \ rel_interior c" by (force simp: rel_interior_eq_empty) then have "x \ c" by (meson subsetD rel_interior_subset) then have "x \ S" using \c facet_of S\ facet_of_imp_subset by blast have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF assms]) have "c \ S" using \c facet_of S\ facet_of_irrefl by blast then have "x \ rel_interior S" by (metis IntI empty_iff \x \ c\ \c \ S\ \c face_of S\ face_of_disjoint_rel_interior) with rels \x \ S\ obtain i where "i \ F" and i: "a i \ x \ b i" by force have "x \ {u. a i \ u \ b i}" by (metis IntD2 InterE \i \ F\ \x \ S\ faceq seq) then have "a i \ x \ b i" by simp then have "a i \ x = b i" using i by auto have "c \ S \ {x. a i \ x = b i}" apply (rule subset_of_face_of [of _ S]) apply (simp add: "*" \i \ F\ facet_of_imp_face_of) apply (simp add: \c face_of S\ face_of_imp_subset) using \a i \ x = b i\ \x \ S\ x by blast then have cface: "c face_of (S \ {x. a i \ x = b i})" by (meson \c face_of S\ face_of_subset inf_le1) have con: "convex (S \ {x. a i \ x = b i})" by (simp add: \convex S\ convex_Int convex_hyperplane) show "\h. h \ F \ c = S \ {x. a h \ x = b h}" apply (rule_tac x=i in exI) apply (simp add: \i \ F\) by (metis (no_types) * \i \ F\ affc facet_of_def less_irrefl face_of_aff_dim_lt [OF con cface]) qed qed lemma face_of_polyhedron_subset_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and c: "c face_of S" and "c \ {}" "c \ S" obtains h where "h \ F" "c \ S \ {x. a h \ x = b h}" proof - have "c \ S" using \c face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" apply (simp add: polyhedron_Int_affine) by (metis \finite F\ faceq seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ {x. a h \ x = b h}) face_of S" if "h \ F" for h apply (rule face_of_Int_supporting_hyperplane_le) using faceq seq that by fastforce have "rel_interior c \ {}" using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain x where "x \ rel_interior c" by auto have rels: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "x \ rel_interior S" by (metis IntI \x \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "x \ S" using \c \ S\ \x \ rel_interior c\ rel_interior_subset by auto then have xint: "x \ \F" using seq by blast have "F \ {}" using assms by (metis affine_Int affine_Inter affine_affine_hull ex_in_conv face_of_affine_trivial) then obtain i where "i \ F" "\ (a i \ x < b i)" using \x \ S\ rels xnot by auto with xint have "a i \ x = b i" by (metis eq_iff mem_Collect_eq not_le Inter_iff faceq) have face: "S \ {x. a i \ x = b i} face_of S" by (simp add: "*" \i \ F\) show ?thesis apply (rule_tac h = i in that) apply (rule \i \ F\) apply (rule subset_of_face_of [OF face \c \ S\]) using \a i \ x = b i\ \x \ rel_interior c\ \x \ S\ apply blast done qed text\Initial part of proof duplicates that above\ proposition face_of_polyhedron_explicit: fixes S :: "'a :: euclidean_space set" assumes "finite F" and seq: "S = affine hull S \ \F" and faceq: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" and psub: "\F'. F' \ F \ S \ affine hull S \ \F'" and c: "c face_of S" and "c \ {}" "c \ S" shows "c = \{S \ {x. a h \ x = b h} | h. h \ F \ c \ S \ {x. a h \ x = b h}}" proof - let ?ab = "\h. {x. a h \ x = b h}" have "c \ S" using \c face_of S\ by (simp add: face_of_imp_subset) have "polyhedron S" apply (simp add: polyhedron_Int_affine) by (metis \finite F\ faceq seq) then have "convex S" by (simp add: polyhedron_imp_convex) then have *: "(S \ ?ab h) face_of S" if "h \ F" for h apply (rule face_of_Int_supporting_hyperplane_le) using faceq seq that by fastforce have "rel_interior c \ {}" using c \c \ {}\ face_of_imp_convex rel_interior_eq_empty by blast then obtain z where z: "z \ rel_interior c" by auto have rels: "rel_interior S = {z \ S. \h\F. a h \ z < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq faceq psub]) then have xnot: "z \ rel_interior S" by (metis IntI \z \ rel_interior c\ c \c \ S\ contra_subsetD empty_iff face_of_disjoint_rel_interior rel_interior_subset) then have "z \ S" using \c \ S\ \z \ rel_interior c\ rel_interior_subset by auto with seq have xint: "z \ \F" by blast have "open (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto simp: \finite F\ open_halfspace_lt open_INT) then obtain e where "0 < e" "ball z e \ (\h\{h \ F. a h \ z < b h}. {w. a h \ w < b h})" by (auto intro: openE [of _ z]) then have e: "\h. \h \ F; a h \ z < b h\ \ ball z e \ {w. a h \ w < b h}" by blast have "c \ (S \ ?ab h) \ z \ S \ ?ab h" if "h \ F" for h proof show "z \ S \ ?ab h \ c \ S \ ?ab h" apply (rule subset_of_face_of [of _ S]) using that \c \ S\ \z \ rel_interior c\ using facet_of_polyhedron_explicit [OF \finite F\ seq faceq psub] unfolding facet_of_def apply auto done next show "c \ S \ ?ab h \ z \ S \ ?ab h" using \z \ rel_interior c\ rel_interior_subset by force qed then have **: "{S \ ?ab h | h. h \ F \ c \ S \ c \ ?ab h} = {S \ ?ab h |h. h \ F \ z \ S \ ?ab h}" by blast have bsub: "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S \ \F \ \{?ab h |h. h \ F \ a h \ z = b h}" if "i \ F" and i: "a i \ z = b i" for i proof - have sub: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ j" if "j \ F" for j proof - have "a j \ z \ b j" using faceq that xint by auto then consider "a j \ z < b j" | "a j \ z = b j" by linarith then have "\G. G \ {?ab h |h. h \ F \ a h \ z = b h} \ ball z e \ G \ j" proof cases assume "a j \ z < b j" then have "ball z e \ {x. a i \ x = b i} \ j" using e [OF \j \ F\] faceq that by (fastforce simp: ball_def) then show ?thesis by (rule_tac x="{x. a i \ x = b i}" in exI) (force simp: \i \ F\ i) next assume eq: "a j \ z = b j" with faceq that show ?thesis by (rule_tac x="{x. a j \ x = b j}" in exI) (fastforce simp add: \j \ F\) qed then show ?thesis by blast qed have 1: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ affine hull S" apply (rule hull_mono) using that \z \ S\ by auto have 2: "affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{?ab h |h. h \ F \ a h \ z = b h}" by (rule hull_minimal) (auto intro: affine_hyperplane) have 3: "ball z e \ \{?ab h |h. h \ F \ a h \ z = b h} \ \F" by (iprover intro: sub Inter_greatest) have *: "\A \ (B :: 'a set); A \ C; E \ C \ D\ \ E \ A \ (B \ D) \ C" for A B C D E by blast show ?thesis by (intro * 1 2 3) qed have "\h. h \ F \ c \ ?ab h" apply (rule face_of_polyhedron_subset_explicit [OF \finite F\ seq faceq psub]) using assms by auto then have fac: "\{S \ ?ab h |h. h \ F \ c \ S \ ?ab h} face_of S" using * by (force simp: \c \ S\ intro: face_of_Inter) have red: "(\a. P a \ T \ S \ \{F x |x. P x}) \ T \ \{S \ F x |x. P x}" for P T F by blast have "ball z e \ affine hull \{S \ ?ab h |h. h \ F \ a h \ z = b h} \ \{S \ ?ab h |h. h \ F \ a h \ z = b h}" apply (rule red) apply (metis seq bsub) done with \0 < e\ have zinrel: "z \ rel_interior (\{S \ ?ab h |h. h \ F \ z \ S \ a h \ z = b h})" by (auto simp: mem_rel_interior_ball \z \ S\) show ?thesis apply (rule face_of_eq [OF c fac]) using z zinrel apply (force simp: **) done qed subsection\More general corollaries from the explicit representation\ corollary facet_of_polyhedron: assumes "polyhedron S" and "c facet_of S" obtains a b where "a \ 0" "S \ {x. a \ x \ b}" "c = S \ {x. a \ x = b}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis obtain i where "i \ F" and c: "c = S \ {x. a i \ x = b i}" using facet_of_polyhedron_explicit [OF \finite F\ seq ab min] assms by force moreover have ssub: "S \ {x. a i \ x \ b i}" apply (subst seq) using \i \ F\ ab by auto ultimately show ?thesis by (rule_tac a = "a i" and b = "b i" in that) (simp_all add: ab) qed corollary face_of_polyhedron: assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" shows "c = \{F. F facet_of S \ c \ F}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis show ?thesis apply (subst face_of_polyhedron_explicit [OF \finite F\ seq ab min]) apply (auto simp: assms facet_of_polyhedron_explicit [OF \finite F\ seq ab min] cong: Collect_cong) done qed lemma face_of_polyhedron_subset_facet: assumes "polyhedron S" and "c face_of S" and "c \ {}" and "c \ S" obtains F where "F facet_of S" "c \ F" using face_of_polyhedron assms by (metis (no_types, lifting) Inf_greatest antisym_conv face_of_imp_subset mem_Collect_eq) lemma exposed_face_of_polyhedron: assumes "polyhedron S" shows "F exposed_face_of S \ F face_of S" proof show "F exposed_face_of S \ F face_of S" by (simp add: exposed_face_of_def) next assume "F face_of S" show "F exposed_face_of S" proof (cases "F = {} \ F = S") case True then show ?thesis using \F face_of S\ exposed_face_of by blast next case False then have "{g. g facet_of S \ F \ g} \ {}" by (metis Collect_empty_eq_bot \F face_of S\ assms empty_def face_of_polyhedron_subset_facet) moreover have "\T. \T facet_of S; F \ T\ \ T exposed_face_of S" by (metis assms exposed_face_of facet_of_imp_face_of facet_of_polyhedron) ultimately have "\{fa. fa facet_of S \ F \ fa} exposed_face_of S" by (metis (no_types, lifting) mem_Collect_eq exposed_face_of_Inter) then show ?thesis using False apply (subst face_of_polyhedron [OF assms \F face_of S\], auto) done qed qed lemma face_of_polyhedron_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" "c face_of S" shows "polyhedron c" by (metis assms face_of_imp_eq_affine_Int polyhedron_Int polyhedron_affine_hull polyhedron_imp_convex) lemma finite_polyhedron_faces: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "finite {F. F face_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have "finite {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" by (simp add: \finite F\) moreover have "{F. F face_of S} - {{}, S} \ {\{S \ {x. a h \ x = b h} |h. h \ F'}| F'. F' \ Pow F}" apply clarify apply (rename_tac c) apply (drule face_of_polyhedron_explicit [OF \finite F\ seq ab min, simplified], simp_all) apply (erule ssubst) apply (rule_tac x="{h \ F. c \ S \ {x. a h \ x = b h}}" in exI, auto) done ultimately show ?thesis by (meson finite.emptyI finite.insertI finite_Diff2 finite_subset) qed lemma finite_polyhedron_exposed_faces: "polyhedron S \ finite {F. F exposed_face_of S}" using exposed_face_of_polyhedron finite_polyhedron_faces by fastforce lemma finite_polyhedron_extreme_points: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {v. v extreme_point_of S}" apply (simp add: face_of_singleton [symmetric]) apply (rule finite_subset [OF _ finite_vimageI [OF finite_polyhedron_faces]], auto) done lemma finite_polyhedron_facets: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ finite {F. F facet_of S}" unfolding facet_of_def by (blast intro: finite_subset [OF _ finite_polyhedron_faces]) proposition rel_interior_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_interior S = S - \{F. F facet_of S}" proof - obtain F where "finite F" and seq: "S = affine hull S \ \F" and faces: "\h. h \ F \ \a b. a \ 0 \ h = {x. a \ x \ b}" and min: "\F'. F' \ F \ S \ (affine hull S) \ \F'" using assms by (simp add: polyhedron_Int_affine_minimal) meson then obtain a b where ab: "\h. h \ F \ a h \ 0 \ h = {x. a h \ x \ b h}" by metis have facet: "(c facet_of S) \ (\h. h \ F \ c = S \ {x. a h \ x = b h})" for c by (rule facet_of_polyhedron_explicit [OF \finite F\ seq ab min]) have rel: "rel_interior S = {x \ S. \h\F. a h \ x < b h}" by (rule rel_interior_polyhedron_explicit [OF \finite F\ seq ab min]) have "a h \ x < b h" if "x \ S" "h \ F" and xnot: "x \ \{F. F facet_of S}" for x h proof - have "x \ \F" using seq that by force with \h \ F\ ab have "a h \ x \ b h" by auto then consider "a h \ x < b h" | "a h \ x = b h" by linarith then show ?thesis proof cases case 1 then show ?thesis . next case 2 have "Collect ((\) x) \ Collect ((\) (\{A. A facet_of S}))" using xnot by fastforce then have "F \ Collect ((\) h)" using 2 \x \ S\ facet by blast with \h \ F\ have "\F \ S \ {x. a h \ x = b h}" by blast with 2 that \x \ \F\ show ?thesis apply simp apply (drule_tac x="\F" in spec) apply (simp add: facet) apply (drule_tac x=h in spec) using seq by auto qed qed moreover have "\h\F. a h \ x \ b h" if "x \ \{F. F facet_of S}" for x using that by (force simp: facet) ultimately show ?thesis by (force simp: rel) qed lemma rel_boundary_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "S - rel_interior S = \ {F. F facet_of S}" using facet_of_imp_subset by (fastforce simp add: rel_interior_of_polyhedron assms) lemma rel_frontier_of_polyhedron: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F facet_of S}" by (simp add: assms rel_frontier_def polyhedron_imp_closed rel_boundary_of_polyhedron) lemma rel_frontier_of_polyhedron_alt: fixes S :: "'a :: euclidean_space set" assumes "polyhedron S" shows "rel_frontier S = \ {F. F face_of S \ (F \ S)}" apply (rule subset_antisym) apply (force simp: rel_frontier_of_polyhedron facet_of_def assms) using face_of_subset_rel_frontier by fastforce text\A characterization of polyhedra as having finitely many faces\ proposition polyhedron_eq_finite_exposed_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F exposed_face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: polyhedron_imp_closed polyhedron_imp_convex finite_polyhedron_exposed_faces) next assume ?rhs then have "closed S" "convex S" and fin: "finite {F. F exposed_face_of S}" by auto show ?lhs proof (cases "S = {}") case True then show ?thesis by auto next case False define F where "F = {h. h exposed_face_of S \ h \ {} \ h \ S}" have "finite F" by (simp add: fin F_def) have hface: "h face_of S" and "\a b. a \ 0 \ S \ {x. a \ x \ b} \ h = S \ {x. a \ x = b}" if "h \ F" for h using exposed_face_of F_def that by simp_all auto then obtain a b where ab: "\h. h \ F \ a h \ 0 \ S \ {x. a h \ x \ b h} \ h = S \ {x. a h \ x = b h}" by metis have *: "False" if paff: "p \ affine hull S" and "p \ S" and pint: "p \ \{{x. a h \ x \ b h} |h. h \ F}" for p proof - have "rel_interior S \ {}" by (simp add: \S \ {}\ \convex S\ rel_interior_eq_empty) then obtain c where c: "c \ rel_interior S" by auto with rel_interior_subset have "c \ S" by blast have ccp: "closed_segment c p \ affine hull S" by (meson affine_affine_hull affine_imp_convex c closed_segment_subset hull_subset paff rel_interior_subset subsetCE) obtain x where xcl: "x \ closed_segment c p" and "x \ S" and xnot: "x \ rel_interior S" using connected_openin [of "closed_segment c p"] apply simp apply (drule_tac x="closed_segment c p \ rel_interior S" in spec) apply (erule impE) apply (force simp: openin_rel_interior openin_Int intro: openin_subtopology_Int_subset [OF _ ccp]) apply (drule_tac x="closed_segment c p \ (- S)" in spec) using rel_interior_subset \closed S\ c \p \ S\ apply blast done then obtain \ where "0 \ \" "\ \ 1" and xeq: "x = (1 - \) *\<^sub>R c + \ *\<^sub>R p" by (auto simp: in_segment) show False proof (cases "\=0 \ \=1") case True with xeq c xnot \x \ S\ \p \ S\ show False by auto next case False then have xos: "x \ open_segment c p" using \x \ S\ c open_segment_def that(2) xcl xnot by auto have xclo: "x \ closure S" using \x \ S\ closure_subset by blast obtain d where "d \ 0" and dle: "\y. y \ closure S \ d \ x \ d \ y" and dless: "\y. y \ rel_interior S \ d \ x < d \ y" by (metis supporting_hyperplane_relative_frontier [OF \convex S\ xclo xnot]) have sex: "S \ {y. d \ y = d \ x} exposed_face_of S" by (simp add: \closed S\ dle exposed_face_of_Int_supporting_hyperplane_ge [OF \convex S\]) have sne: "S \ {y. d \ y = d \ x} \ {}" using \x \ S\ by blast have sns: "S \ {y. d \ y = d \ x} \ S" by (metis (mono_tags) Int_Collect c subsetD dless not_le order_refl rel_interior_subset) obtain h where "h \ F" "x \ h" apply (rule_tac h="S \ {y. d \ y = d \ x}" in that) apply (simp_all add: F_def sex sne sns \x \ S\) done have abface: "{y. a h \ y = b h} face_of {y. a h \ y \ b h}" using hyperplane_face_of_halfspace_le by blast then have "c \ h" using face_ofD [OF abface xos] \c \ S\ \h \ F\ ab pint \x \ h\ by blast with c have "h \ rel_interior S \ {}" by blast then show False using \h \ F\ F_def face_of_disjoint_rel_interior hface by auto qed qed have "S \ affine hull S \ \{{x. a h \ x \ b h} |h. h \ F}" using ab by (auto simp: hull_subset) moreover have "affine hull S \ \{{x. a h \ x \ b h} |h. h \ F} \ S" using * by blast ultimately have "S = affine hull S \ \ {{x. a h \ x \ b h} |h. h \ F}" .. then show ?thesis apply (rule ssubst) apply (force intro: polyhedron_affine_hull polyhedron_halfspace_le simp: \finite F\) done qed qed corollary polyhedron_eq_finite_faces: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ closed S \ convex S \ finite {F. F face_of S}" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polyhedron_faces polyhedron_imp_closed polyhedron_imp_convex) next assume ?rhs then show ?lhs by (force simp: polyhedron_eq_finite_exposed_faces exposed_face_of intro: finite_subset) qed lemma polyhedron_linear_image_eq: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "linear h" "bij h" shows "polyhedron (h ` S) \ polyhedron S" proof - have *: "{f. P f} = (image h) ` {f. P (h ` f)}" for P apply safe apply (rule_tac x="inv h ` x" in image_eqI) apply (auto simp: \bij h\ bij_is_surj image_f_inv_f) done have "inj h" using bij_is_inj assms by blast then have injim: "inj_on ((`) h) A" for A by (simp add: inj_on_def inj_image_eq_iff) show ?thesis using \linear h\ \inj h\ apply (simp add: polyhedron_eq_finite_faces closed_injective_linear_image_eq) apply (simp add: * face_of_linear_image [of h _ S, symmetric] finite_image_iff injim) done qed lemma polyhedron_negations: fixes S :: "'a :: euclidean_space set" shows "polyhedron S \ polyhedron(image uminus S)" by (subst polyhedron_linear_image_eq) (auto simp: bij_uminus intro!: linear_uminus) subsection\Relation between polytopes and polyhedra\ proposition polytope_eq_bounded_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S \ bounded S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: finite_polytope_faces polyhedron_eq_finite_faces polytope_imp_closed polytope_imp_convex polytope_imp_bounded) next assume ?rhs then show ?lhs unfolding polytope_def apply (rule_tac x="{v. v extreme_point_of S}" in exI) apply (simp add: finite_polyhedron_extreme_points Krein_Milman_Minkowski compact_eq_bounded_closed polyhedron_imp_closed polyhedron_imp_convex) done qed lemma polytope_Int: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polytope T\ \ polytope(S \ T)" by (simp add: polytope_eq_bounded_polyhedron bounded_Int) lemma polytope_Int_polyhedron: fixes S :: "'a :: euclidean_space set" shows "\polytope S; polyhedron T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polyhedron_Int_polytope: fixes S :: "'a :: euclidean_space set" shows "\polyhedron S; polytope T\ \ polytope(S \ T)" by (simp add: bounded_Int polytope_eq_bounded_polyhedron) lemma polytope_imp_polyhedron: fixes S :: "'a :: euclidean_space set" shows "polytope S \ polyhedron S" by (simp add: polytope_eq_bounded_polyhedron) lemma polytope_facet_exists: fixes p :: "'a :: euclidean_space set" assumes "polytope p" "0 < aff_dim p" obtains F where "F facet_of p" proof (cases "p = {}") case True with assms show ?thesis by auto next case False then obtain v where "v extreme_point_of p" using extreme_point_exists_convex by (blast intro: \polytope p\ polytope_imp_compact polytope_imp_convex) then show ?thesis by (metis face_of_polyhedron_subset_facet polytope_imp_polyhedron aff_dim_sing all_not_in_conv assms face_of_singleton less_irrefl singletonI that) qed lemma polyhedron_interval [iff]: "polyhedron(cbox a b)" by (metis polytope_imp_polyhedron polytope_interval) lemma polyhedron_convex_hull: fixes S :: "'a :: euclidean_space set" shows "finite S \ polyhedron(convex hull S)" by (simp add: polytope_convex_hull polytope_imp_polyhedron) subsection\Relative and absolute frontier of a polytope\ lemma rel_boundary_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "\ affine_dependent S" shows "(convex hull S) - rel_interior(convex hull S) = (\a\S. convex hull (S - {a}))" proof - have "finite S" by (metis assms aff_independent_finite) then consider "card S = 0" | "card S = 1" | "2 \ card S" by arith then show ?thesis proof cases case 1 then have "S = {}" by (simp add: \finite S\) then show ?thesis by simp next case 2 show ?thesis by (auto intro: card_1_singletonE [OF \card S = 1\]) next case 3 with assms show ?thesis by (auto simp: polyhedron_convex_hull rel_boundary_of_polyhedron facet_of_convex_hull_affine_independent_alt \finite S\) qed qed proposition frontier_of_convex_hull: fixes S :: "'a::euclidean_space set" assumes "card S = Suc (DIM('a))" shows "frontier(convex hull S) = \ {convex hull (S - {a}) | a. a \ S}" proof (cases "affine_dependent S") case True have [iff]: "finite S" using assms using card_infinite by force then have ccs: "closed (convex hull S)" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) { fix x T assume "finite T" "T \ S" "int (card T) \ aff_dim S + 1" "x \ convex hull T" then have "S \ T" using True \finite S\ aff_dim_le_card affine_independent_iff_card by fastforce then obtain a where "a \ S" "a \ T" using \T \ S\ by blast then have "x \ (\a\S. convex hull (S - {a}))" using True affine_independent_iff_card [of S] apply simp apply (metis (no_types, hide_lams) Diff_eq_empty_iff Diff_insert0 \a \ T\ \T \ S\ \x \ convex hull T\ hull_mono insert_Diff_single subsetCE) done } note * = this have 1: "convex hull S \ (\ a\S. convex hull (S - {a}))" apply (subst caratheodory_aff_dim) apply (blast intro: *) done have 2: "\((\a. convex hull (S - {a})) ` S) \ convex hull S" by (rule Union_least) (metis (no_types, lifting) Diff_subset hull_mono imageE) show ?thesis using True apply (simp add: segment_convex_hull frontier_def) using interior_convex_hull_eq_empty [OF assms] apply (simp add: closure_closed [OF ccs]) apply (rule subset_antisym) using 1 apply blast using 2 apply blast done next case False then have "frontier (convex hull S) = (convex hull S) - rel_interior(convex hull S)" apply (simp add: rel_boundary_of_convex_hull [symmetric] frontier_def) by (metis aff_independent_finite assms closure_convex_hull finite_imp_compact_convex_hull hull_hull interior_convex_hull_eq_empty rel_interior_nonempty_interior) also have "... = \{convex hull (S - {a}) |a. a \ S}" proof - have "convex hull S - rel_interior (convex hull S) = rel_frontier (convex hull S)" by (simp add: False aff_independent_finite polyhedron_convex_hull rel_boundary_of_polyhedron rel_frontier_of_polyhedron) then show ?thesis by (simp add: False rel_frontier_convex_hull_cases) qed finally show ?thesis . qed subsection\Special case of a triangle\ proposition frontier_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "frontier(convex hull {a,b,c}) = closed_segment a b \ closed_segment b c \ closed_segment c a" (is "?lhs = ?rhs") proof (cases "b = a \ c = a \ c = b") case True then show ?thesis by (auto simp: assms segment_convex_hull frontier_def empty_interior_convex_hull insert_commute card_insert_le_m1 hull_inc insert_absorb) next case False then have [simp]: "card {a, b, c} = Suc (DIM('a))" by (simp add: card_insert Set.insert_Diff_if assms) show ?thesis proof show "?lhs \ ?rhs" using False by (force simp: segment_convex_hull frontier_of_convex_hull insert_Diff_if insert_commute split: if_split_asm) show "?rhs \ ?lhs" using False apply (simp add: frontier_of_convex_hull segment_convex_hull) apply (intro conjI subsetI) apply (rule_tac X="convex hull {a,b}" in UnionI; force simp: Set.insert_Diff_if) apply (rule_tac X="convex hull {b,c}" in UnionI; force) apply (rule_tac X="convex hull {a,c}" in UnionI; force simp: insert_commute Set.insert_Diff_if) done qed qed corollary inside_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "inside (closed_segment a b \ closed_segment b c \ closed_segment c a) = interior(convex hull {a,b,c})" by (metis assms frontier_of_triangle bounded_empty bounded_insert convex_convex_hull inside_frontier_eq_interior bounded_convex_hull) corollary interior_of_triangle: fixes a :: "'a::euclidean_space" assumes "DIM('a) = 2" shows "interior(convex hull {a,b,c}) = convex hull {a,b,c} - (closed_segment a b \ closed_segment b c \ closed_segment c a)" using interior_subset by (force simp: frontier_of_triangle [OF assms, symmetric] frontier_def Diff_Diff_Int) subsection\Subdividing a cell complex\ lemma subdivide_interval: fixes x::real assumes "a < \x - y\" "0 < a" obtains n where "n \ \" "x < n * a \ n * a < y \ y < n * a \ n * a < x" proof - consider "a + x < y" | "a + y < x" using assms by linarith then show ?thesis proof cases case 1 let ?n = "of_int (floor (x/a)) + 1" have x: "x < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + x" apply (simp add: algebra_simps) by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) also have "... < y" by (rule 1) finally have "?n * a < y" . with x show ?thesis using Ints_1 Ints_add Ints_of_int that by blast next case 2 let ?n = "of_int (floor (y/a)) + 1" have y: "y < ?n * a" by (meson \0 < a\ divide_less_eq floor_eq_iff) have "?n * a \ a + y" apply (simp add: algebra_simps) by (metis \0 < a\ floor_correct less_irrefl nonzero_mult_div_cancel_left real_mult_le_cancel_iff2 times_divide_eq_right) also have "... < x" by (rule 2) finally have "?n * a < x" . then show ?thesis using Ints_1 Ints_add Ints_of_int that y by blast qed qed lemma cell_subdivision_lemma: assumes "finite \" and "\X. X \ \ \ polytope X" and "\X. X \ \ \ aff_dim X \ d" and "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and "finite I" shows "\\. \\ = \\ \ finite \ \ (\C \ \. \D. D \ \ \ C \ D) \ (\C \ \. \x \ C. \D. D \ \ \ x \ D \ D \ C) \ (\X \ \. polytope X) \ (\X \ \. aff_dim X \ d) \ (\X \ \. \Y \ \. X \ Y face_of X \ X \ Y face_of Y) \ (\X \ \. \x \ X. \y \ X. \a b. (a,b) \ I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b)" using \finite I\ proof induction case empty then show ?case by (rule_tac x="\" in exI) (auto simp: assms) next case (insert ab I) then obtain \ where eq: "\\ = \\" and "finite \" and sub1: "\C. C \ \ \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" and I: "\X x y a b. \X \ \; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" by (auto simp: that) obtain a b where "ab = (a,b)" by fastforce let ?\ = "(\X. X \ {x. a \ x \ b}) ` \ \ (\X. X \ {x. a \ x \ b}) ` \" have eqInt: "(S \ Collect P) \ (T \ Collect Q) = (S \ T) \ (Collect P \ Collect Q)" for S T::"'a set" and P Q by blast show ?case proof (intro conjI exI) show "\?\ = \\" by (force simp: eq [symmetric]) show "finite ?\" using \finite \\ by force show "\X \ ?\. polytope X" by (force simp: poly polytope_Int_polyhedron polyhedron_halfspace_le polyhedron_halfspace_ge) show "\X \ ?\. aff_dim X \ d" by (auto; metis order_trans aff aff_dim_subset inf_le1) show "\X \ ?\. \x \ X. \y \ X. \a b. (a,b) \ insert ab I \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" using \ab = (a, b)\ I by fastforce show "\X \ ?\. \Y \ ?\. X \ Y face_of X \ X \ Y face_of Y" by (auto simp: eqInt halfspace_Int_eq face_of_Int_Int face face_of_halfspace_le face_of_halfspace_ge) show "\C \ ?\. \D. D \ \ \ C \ D" using sub1 by force show "\C\\. \x\C. \D. D \ ?\ \ x \ D \ D \ C" proof (intro ballI) fix C z assume "C \ \" "z \ C" with sub2 obtain D where D: "D \ \" "z \ D" "D \ C" by blast have "D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C \ D \ \ \ z \ D \ {x. a \ x \ b} \ D \ {x. a \ x \ b} \ C" using linorder_class.linear [of "a \ z" b] D by blast then show "\D. D \ ?\ \ z \ D \ D \ C" by blast qed qed qed proposition cell_complex_subdivision_exists: fixes \ :: "'a::euclidean_space set set" assumes "0 < e" "finite \" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ d" and face: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" obtains "\'" where "finite \'" "\\' = \\" "\X. X \ \' \ diameter X < e" "\X. X \ \' \ polytope X" "\X. X \ \' \ aff_dim X \ d" "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X \ X \ Y face_of Y" "\C. C \ \' \ \D. D \ \ \ C \ D" "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" proof - have "bounded(\\)" by (simp add: \finite \\ poly bounded_Union polytope_imp_bounded) then obtain B where "B > 0" and B: "\x. x \ \\ \ norm x < B" by (meson bounded_pos_less) define C where "C \ {z \ \. \z * e / 2 / real DIM('a)\ \ B}" define I where "I \ \i \ Basis. \j \ C. { (i::'a, j * e / 2 / DIM('a)) }" have "finite C" using finite_int_segment [of "-B / (e / 2 / DIM('a))" "B / (e / 2 / DIM('a))"] apply (simp add: C_def) apply (erule rev_finite_subset) using \0 < e\ apply (auto simp: field_split_simps) done then have "finite I" by (simp add: I_def) obtain \' where eq: "\\' = \\" and "finite \'" and poly: "\X. X \ \' \ polytope X" and aff: "\X. X \ \' \ aff_dim X \ d" and face: "\X Y. \X \ \'; Y \ \'\ \ X \ Y face_of X \ X \ Y face_of Y" and I: "\X x y a b. \X \ \'; x \ X; y \ X; (a,b) \ I\ \ a \ x \ b \ a \ y \ b \ a \ x \ b \ a \ y \ b" and sub1: "\C. C \ \' \ \D. D \ \ \ C \ D" and sub2: "\C x. C \ \ \ x \ C \ \D. D \ \' \ x \ D \ D \ C" apply (rule exE [OF cell_subdivision_lemma]) using assms \finite I\ apply auto done show ?thesis proof (rule_tac \'="\'" in that) show "diameter X < e" if "X \ \'" for X proof - have "diameter X \ e/2" proof (rule diameter_le) show "norm (x - y) \ e / 2" if "x \ X" "y \ X" for x y proof - have "norm x < B" "norm y < B" using B \X \ \'\ eq that by fastforce+ have "norm (x - y) \ (\b\Basis. \(x-y) \ b\)" by (rule norm_le_l1) also have "... \ of_nat (DIM('a)) * (e / 2 / DIM('a))" proof (rule sum_bounded_above) fix i::'a assume "i \ Basis" then have I': "\z b. \z \ C; b = z * e / (2 * real DIM('a))\ \ i \ x \ b \ i \ y \ b \ i \ x \ b \ i \ y \ b" using I \X \ \'\ that by (fastforce simp: I_def) show "\(x - y) \ i\ \ e / 2 / real DIM('a)" proof (rule ccontr) assume "\ \(x - y) \ i\ \ e / 2 / real DIM('a)" then have xyi: "\i \ x - i \ y\ > e / 2 / real DIM('a)" by (simp add: inner_commute inner_diff_right) obtain n where "n \ \" and n: "i \ x < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ y \ i \ y < n * (e / 2 / real DIM('a)) \ n * (e / 2 / real DIM('a)) < i \ x" using subdivide_interval [OF xyi] DIM_positive \0 < e\ by (auto simp: zero_less_divide_iff) have "\i \ x\ < B" by (metis \i \ Basis\ \norm x < B\ inner_commute norm_bound_Basis_lt) have "\i \ y\ < B" by (metis \i \ Basis\ \norm y < B\ inner_commute norm_bound_Basis_lt) have *: "\n * e\ \ B * (2 * real DIM('a))" if "\ix\ < B" "\iy\ < B" and ix: "ix * (2 * real DIM('a)) < n * e" and iy: "n * e < iy * (2 * real DIM('a))" for ix iy proof (rule abs_leI) have "iy * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\iy\ < B\ in linarith)+ then show "n * e \ B * (2 * real DIM('a))" using iy by linarith next have "- ix * (2 * real DIM('a)) \ B * (2 * real DIM('a))" by (rule mult_right_mono) (use \\ix\ < B\ in linarith)+ then show "- (n * e) \ B * (2 * real DIM('a))" using ix by linarith qed have "n \ C" using \n \ \\ n by (auto simp: C_def divide_simps intro: * \\i \ x\ < B\ \\i \ y\ < B\) show False using I' [OF \n \ C\ refl] n by auto qed qed also have "... = e / 2" by simp finally show ?thesis . qed qed (use \0 < e\ in force) also have "... < e" by (simp add: \0 < e\) finally show ?thesis . qed qed (auto simp: eq poly aff face sub1 sub2 \finite \'\) qed subsection\Simplexes\ text\The notion of n-simplex for integer \<^term>\n \ -1\\ definition\<^marker>\tag important\ simplex :: "int \ 'a::euclidean_space set \ bool" (infix "simplex" 50) where "n simplex S \ \C. \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C" lemma simplex: "n simplex S \ (\C. finite C \ \ affine_dependent C \ int(card C) = n + 1 \ S = convex hull C)" by (auto simp add: simplex_def intro: aff_independent_finite) lemma simplex_convex_hull: "\ affine_dependent C \ int(card C) = n + 1 \ n simplex (convex hull C)" by (auto simp add: simplex_def) lemma convex_simplex: "n simplex S \ convex S" by (metis convex_convex_hull simplex_def) lemma compact_simplex: "n simplex S \ compact S" unfolding simplex using finite_imp_compact_convex_hull by blast lemma closed_simplex: "n simplex S \ closed S" by (simp add: compact_imp_closed compact_simplex) lemma simplex_imp_polytope: "n simplex S \ polytope S" unfolding simplex_def polytope_def using aff_independent_finite by blast lemma simplex_imp_polyhedron: "n simplex S \ polyhedron S" by (simp add: polytope_imp_polyhedron simplex_imp_polytope) lemma simplex_dim_ge: "n simplex S \ -1 \ n" by (metis (no_types, hide_lams) aff_dim_geq affine_independent_iff_card diff_add_cancel diff_diff_eq2 simplex_def) lemma simplex_empty [simp]: "n simplex {} \ n = -1" proof assume "n simplex {}" then show "n = -1" unfolding simplex by (metis card_empty convex_hull_eq_empty diff_0 diff_eq_eq of_nat_0) next assume "n = -1" then show "n simplex {}" by (fastforce simp: simplex) qed lemma simplex_minus_1 [simp]: "-1 simplex S \ S = {}" by (metis simplex cancel_comm_monoid_add_class.diff_cancel card_0_eq diff_minus_eq_add of_nat_eq_0_iff simplex_empty) lemma aff_dim_simplex: "n simplex S \ aff_dim S = n" by (metis simplex add.commute add_diff_cancel_left' aff_dim_convex_hull affine_independent_iff_card) lemma zero_simplex_sing: "0 simplex {a}" apply (simp add: simplex_def) by (metis affine_independent_1 card_empty card_insert_disjoint convex_hull_singleton empty_iff finite.emptyI) lemma simplex_sing [simp]: "n simplex {a} \ n = 0" using aff_dim_simplex aff_dim_sing zero_simplex_sing by blast lemma simplex_zero: "0 simplex S \ (\a. S = {a})" apply (auto simp: ) using aff_dim_eq_0 aff_dim_simplex by blast lemma one_simplex_segment: "a \ b \ 1 simplex closed_segment a b" apply (simp add: simplex_def) apply (rule_tac x="{a,b}" in exI) apply (auto simp: segment_convex_hull) done lemma simplex_segment_cases: "(if a = b then 0 else 1) simplex closed_segment a b" by (auto simp: one_simplex_segment) lemma simplex_segment: "\n. n simplex closed_segment a b" using simplex_segment_cases by metis lemma polytope_lowdim_imp_simplex: assumes "polytope P" "aff_dim P \ 1" obtains n where "n simplex P" proof (cases "P = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis by (metis assms compact_convex_collinear_segment collinear_aff_dim polytope_imp_compact polytope_imp_convex simplex_segment_cases that) qed lemma simplex_insert_dimplus1: fixes n::int assumes "n simplex S" and a: "a \ affine hull S" shows "(n+1) simplex (convex hull (insert a S))" proof - obtain C where C: "finite C" "\ affine_dependent C" "int(card C) = n+1" and S: "S = convex hull C" using assms unfolding simplex by force show ?thesis unfolding simplex proof (intro exI conjI) have "aff_dim S = n" using aff_dim_simplex assms(1) by blast moreover have "a \ affine hull C" using S a affine_hull_convex_hull by blast moreover have "a \ C" using S a hull_inc by fastforce ultimately show "\ affine_dependent (insert a C)" by (simp add: C S aff_dim_convex_hull aff_dim_insert affine_independent_iff_card) next have "a \ C" using S a hull_inc by fastforce then show "int (card (insert a C)) = n + 1 + 1" by (simp add: C) next show "convex hull insert a S = convex hull (insert a C)" by (simp add: S convex_hull_insert_segments) qed (use C in auto) qed subsection \Simplicial complexes and triangulations\ definition\<^marker>\tag important\ simplicial_complex where "simplicial_complex \ \ finite \ \ (\S \ \. \n. n simplex S) \ (\F S. S \ \ \ F face_of S \ F \ \) \ (\S S'. S \ \ \ S' \ \ \ (S \ S') face_of S \ (S \ S') face_of S')" definition\<^marker>\tag important\ triangulation where "triangulation \ \ finite \ \ (\T \ \. \n. n simplex T) \ (\T T'. T \ \ \ T' \ \ \ (T \ T') face_of T \ (T \ T') face_of T')" subsection\Refining a cell complex to a simplicial complex\ proposition convex_hull_insert_Int_eq: fixes z :: "'a :: euclidean_space" assumes z: "z \ rel_interior S" and T: "T \ rel_frontier S" and U: "U \ rel_frontier S" and "convex S" "convex T" "convex U" shows "convex hull (insert z T) \ convex hull (insert z U) = convex hull (insert z (T \ U))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof (cases "T={} \ U={}") case True then show ?thesis by auto next case False then have "T \ {}" "U \ {}" by auto have TU: "convex (T \ U)" by (simp add: \convex T\ \convex U\ convex_Int) have "(\x\T. closed_segment z x) \ (\x\U. closed_segment z x) \ (if T \ U = {} then {z} else \((closed_segment z) ` (T \ U)))" (is "_ \ ?IF") proof clarify fix x t u assume xt: "x \ closed_segment z t" and xu: "x \ closed_segment z u" and "t \ T" "u \ U" then have ne: "t \ z" "u \ z" using T U z unfolding rel_frontier_def by blast+ show "x \ ?IF" proof (cases "x = z") case True then show ?thesis by auto next case False have t: "t \ closure S" using T \t \ T\ rel_frontier_def by auto have u: "u \ closure S" using U \u \ U\ rel_frontier_def by auto show ?thesis proof (cases "t = u") case True then show ?thesis using \t \ T\ \u \ U\ xt by auto next case False have tnot: "t \ closed_segment u z" proof - have "t \ closure S - rel_interior S" using T \t \ T\ rel_frontier_def by blast then have "t \ open_segment z u" by (meson DiffD2 rel_interior_closure_convex_segment [OF \convex S\ z u] subsetD) then show ?thesis by (simp add: \t \ u\ \t \ z\ open_segment_commute open_segment_def) qed moreover have "u \ closed_segment z t" using rel_interior_closure_convex_segment [OF \convex S\ z t] \u \ U\ \u \ z\ U [unfolded rel_frontier_def] tnot by (auto simp: closed_segment_eq_open) ultimately have "\(between (t,u) z | between (u,z) t | between (z,t) u)" if "x \ z" using that xt xu apply (simp add: between_mem_segment [symmetric]) by (metis between_commute between_trans_2 between_antisym) then have "\ collinear {t, z, u}" if "x \ z" by (auto simp: that collinear_between_cases between_commute) moreover have "collinear {t, z, x}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xt) moreover have "collinear {z, x, u}" by (metis closed_segment_commute collinear_2 collinear_closed_segment collinear_triples ends_in_segment(1) insert_absorb insert_absorb2 xu) ultimately have False using collinear_3_trans [of t z x u] \x \ z\ by blast then show ?thesis by metis qed qed qed then show ?thesis using False \convex T\ \convex U\ TU by (simp add: convex_hull_insert_segments hull_same split: if_split_asm) qed show "?rhs \ ?lhs" by (metis inf_greatest hull_mono inf.cobounded1 inf.cobounded2 insert_mono) qed lemma simplicial_subdivision_aux: assumes "finite \" and "\C. C \ \ \ polytope C" and "\C. C \ \ \ aff_dim C \ of_nat n" and "\C F. \C \ \; F face_of C\ \ F \ \" and "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" shows "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ \\ = \\ \ (\C \ \. \F. finite F \ F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" using assms proof (induction n arbitrary: \ rule: less_induct) case (less n) then have poly\: "\C. C \ \ \ polytope C" and aff\: "\C. C \ \ \ aff_dim C \ of_nat n" and face\: "\C F. \C \ \; F face_of C\ \ F \ \" and intface\: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" by metis+ show ?case proof (cases "n \ 1") case True have "\s. \n \ 1; s \ \\ \ \m. m simplex s" using poly\ aff\ by (force intro: polytope_lowdim_imp_simplex) then show ?thesis unfolding simplicial_complex_def apply (rule_tac x="\" in exI) using True by (auto simp: less.prems) next case False define \ where "\ \ {C \ \. aff_dim C < n}" have "finite \" "\C. C \ \ \ polytope C" "\C. C \ \ \ aff_dim C \ int (n - 1)" "\C F. \C \ \; F face_of C\ \ F \ \" "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" using less.prems apply (auto simp: \_def) by (metis aff_dim_subset face_of_imp_subset less_le not_le) with less.IH [of "n-1" \] False obtain \ where "simplicial_complex \" and aff_dim\: "\K. K \ \ \ aff_dim K \ int (n - 1)" and "\\ = \\" and fin\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and C\: "\K. K \ \ \ \C. C \ \ \ K \ C" by auto then have "finite \" and simpl\: "\S. S \ \ \ \n. n simplex S" and face\: "\F S. \S \ \; F face_of S\ \ F \ \" and faceI\: "\S S'. \S \ \; S' \ \\ \ (S \ S') face_of S \ (S \ S') face_of S'" by (auto simp: simplicial_complex_def) define \ where "\ \ {C \ \. aff_dim C = n}" have "finite \" by (simp add: \_def less.prems(1)) have poly\: "\C. C \ \ \ polytope C" and convex\: "\C. C \ \ \ convex C" and closed\: "\C. C \ \ \ closed C" by (auto simp: \_def poly\ polytope_imp_convex polytope_imp_closed) have in_rel_interior: "(SOME z. z \ rel_interior C) \ rel_interior C" if "C \ \" for C using that poly\ polytope_imp_convex rel_interior_aff_dim some_in_eq by (fastforce simp: \_def) have *: "\T. \ affine_dependent T \ card T \ n \ aff_dim K < n \ K = convex hull T" if "K \ \" for K proof - obtain r where r: "r simplex K" using \K \ \\ simpl\ by blast have "r = aff_dim K" using \r simplex K\ aff_dim_simplex by blast with r show ?thesis unfolding simplex_def using False \\K. K \ \ \ aff_dim K \ int (n - 1)\ that by fastforce qed have ahK_C_disjoint: "affine hull K \ rel_interior C = {}" if "C \ \" "K \ \" "K \ rel_frontier C" for C K proof - have "convex C" "closed C" by (auto simp: convex\ closed\ \C \ \\) obtain F where F: "F face_of C" and "F \ C" "K \ F" proof - obtain L where "L \ \" "K \ L" using \K \ \\ C\ by blast have "K \ rel_frontier C" by (simp add: \K \ rel_frontier C\) also have "... \ C" by (simp add: \closed C\ rel_frontier_def subset_iff) finally have "K \ C" . have "L \ C face_of C" using \_def \_def \C \ \\ \L \ \\ intface\ by auto moreover have "L \ C \ C" using \C \ \\ \L \ \\ apply (clarsimp simp: \_def \_def) by (metis aff_dim_subset inf_le1 not_le) moreover have "K \ L \ C" using \C \ \\ \L \ \\ \K \ C\ \K \ L\ by (auto simp: \_def \_def) ultimately show ?thesis using that by metis qed have "affine hull F \ rel_interior C = {}" by (rule affine_hull_face_of_disjoint_rel_interior [OF \convex C\ F \F \ C\]) with hull_mono [OF \K \ F\] show "affine hull K \ rel_interior C = {}" by fastforce qed let ?\ = "(\C \ \. \K \ \ \ Pow (rel_frontier C). {convex hull (insert (SOME z. z \ rel_interior C) K)})" have "\\. simplicial_complex \ \ (\K \ \. aff_dim K \ of_nat n) \ (\C \ \. \F. F \ \ \ C = \F) \ (\K \ \. \C. C \ \ \ K \ C)" proof (rule exI, intro conjI ballI) show "simplicial_complex (\ \ ?\)" unfolding simplicial_complex_def proof (intro conjI impI ballI allI) show "finite (\ \ ?\)" using \finite \\ \finite \\ by simp show "\n. n simplex S" if "S \ \ \ ?\" for S using that ahK_C_disjoint in_rel_interior simpl\ simplex_insert_dimplus1 by fastforce show "F \ \ \ ?\" if S: "S \ \ \ ?\ \ F face_of S" for F S proof - have "F \ \" if "S \ \" using S face\ that by blast moreover have "F \ \ \ ?\" if "F face_of S" "C \ \" "K \ \" and "K \ rel_frontier C" and S: "S = convex hull insert (SOME z. z \ rel_interior C) K" for C K proof - let ?z = "SOME z. z \ rel_interior C" have "?z \ rel_interior C" by (simp add: in_rel_interior \C \ \\) moreover obtain I where "\ affine_dependent I" "card I \ n" "aff_dim K < int n" "K = convex hull I" using * [OF \K \ \\] by auto ultimately have "?z \ affine hull I" using ahK_C_disjoint affine_hull_convex_hull that by blast have "compact I" "finite I" by (auto simp: \\ affine_dependent I\ aff_independent_finite finite_imp_compact) moreover have "F face_of convex hull insert ?z I" by (metis S \F face_of S\ \K = convex hull I\ convex_hull_eq_empty convex_hull_insert_segments hull_hull) ultimately obtain J where "J \ insert ?z I" "F = convex hull J" using face_of_convex_hull_subset [of "insert ?z I" F] by auto show ?thesis proof (cases "?z \ J") case True have "F \ (\K\\ \ Pow (rel_frontier C). {convex hull insert ?z K})" proof have "convex hull (J - {?z}) face_of K" by (metis True \J \ insert ?z I\ \K = convex hull I\ \\ affine_dependent I\ face_of_convex_hull_affine_independent subset_insert_iff) then have "convex hull (J - {?z}) \ \" by (rule face\ [OF \K \ \\]) moreover have "\x. x \ convex hull (J - {?z}) \ x \ rel_frontier C" by (metis True \J \ insert ?z I\ \K = convex hull I\ subsetD hull_mono subset_insert_iff that(4)) ultimately show "convex hull (J - {?z}) \ \ \ Pow (rel_frontier C)" by auto let ?F = "convex hull insert ?z (convex hull (J - {?z}))" have "F \ ?F" apply (clarsimp simp: \F = convex hull J\) by (metis True subsetD hull_mono hull_subset subset_insert_iff) moreover have "?F \ F" apply (clarsimp simp: \F = convex hull J\) by (metis (no_types, lifting) True convex_hull_eq_empty convex_hull_insert_segments hull_hull insert_Diff) ultimately show "F \ {?F}" by auto qed with \C\\\ show ?thesis by auto next case False then have "F \ \" using face_of_convex_hull_affine_independent [OF \\ affine_dependent I\] by (metis Int_absorb2 Int_insert_right_if0 \F = convex hull J\ \J \ insert ?z I\ \K = convex hull I\ face\ inf_le2 \K \ \\) then show "F \ \ \ ?\" by blast qed qed ultimately show ?thesis using that by auto qed have "(S \ S' face_of S) \ (S \ S' face_of S')" if "S \ \ \ ?\" "S' \ \ \ ?\" for S S' proof - have symmy: "\\X Y. R X Y \ R Y X; \X Y. \X \ \; Y \ \\ \ R X Y; \X Y. \X \ \; Y \ ?\\ \ R X Y; \X Y. \X \ ?\; Y \ ?\\ \ R X Y\ \ R S S'" for R using that by (metis (no_types, lifting) Un_iff) show ?thesis proof (rule symmy) show "Y \ X face_of Y \ Y \ X face_of X" if "X \ Y face_of X \ X \ Y face_of Y" for X Y :: "'a set" by (simp add: inf_commute that) next show "X \ Y face_of X \ X \ Y face_of Y" if "X \ \" and "Y \ \" for X Y by (simp add: faceI\ that) next show "X \ Y face_of X \ X \ Y face_of Y" if XY: "X \ \" "Y \ ?\" for X Y proof - obtain C K where "C \ \" "K \ \" "K \ rel_frontier C" and Y: "Y = convex hull insert (SOME z. z \ rel_interior C) K" using XY by blast have "convex C" by (simp add: \C \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_iff) let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast obtain D where "D \ \" "X \ D" using C\ \X \ \\ by blast have "D \ rel_interior C = (C \ D) \ rel_interior C" using rel_interior_subset by blast also have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \_def \C \ \\ \D \ \\ intface\ by blast show "C \ D \ C" by (metis (mono_tags, lifting) Int_lower2 \_def \_def \C \ \\ \D \ \\ aff_dim_subset mem_Collect_eq not_le) qed finally have DC: "D \ rel_interior C = {}" . have eq: "X \ convex hull (insert ?z K) = X \ convex hull K" apply (rule Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) using DC by (meson \X \ D\ disjnt_def disjnt_subset1) obtain I where I: "\ affine_dependent I" and Keq: "K = convex hull I" and [simp]: "convex hull K = K" using "*" \K \ \\ by force then have "?z \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull z by blast have "X \ K face_of K" by (simp add: \K \ \\ faceI\ \X \ \\) also have "... face_of convex hull insert ?z K" by (metis I Keq \?z \ affine hull I\ aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert) finally have "X \ K face_of convex hull insert ?z K" . then show ?thesis using "*" \K \ \\ faceI\ that(1) by (fastforce simp add: Y eq) qed next show "X \ Y face_of X \ X \ Y face_of Y" if XY: "X \ ?\" "Y \ ?\" for X Y proof - obtain C K D L where "C \ \" "K \ \" "K \ rel_frontier C" and X: "X = convex hull insert (SOME z. z \ rel_interior C) K" and "D \ \" "L \ \" "L \ rel_frontier D" and Y: "Y = convex hull insert (SOME z. z \ rel_interior D) L" using XY by blast let ?z = "(SOME z. z \ rel_interior C)" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast have "convex C" by (simp add: \C \ \\ convex\) have "convex K" using "*" \K \ \\ by blast have "convex L" by (meson \L \ \\ convex_simplex simpl\) show ?thesis proof (cases "D=C") case True then have "L \ rel_frontier C" using \L \ rel_frontier D\ by auto show ?thesis apply (simp add: X Y True) apply (simp add: convex_hull_insert_Int_eq [OF z] \K \ rel_frontier C\ \L \ rel_frontier C\ \convex C\ \convex K\ \convex L\) using face_of_polytope_insert2 by (metis "*" IntI \C \ \\ \K \ \\ \L \ \\\K \ rel_frontier C\ \L \ rel_frontier C\ aff_independent_finite ahK_C_disjoint empty_iff faceI\ polytope_convex_hull z) next case False have "convex D" by (simp add: \D \ \\ convex\) have "K \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed rel_frontier_def subset_eq) have "L \ D" by (metis DiffE \D \ \\ \L \ rel_frontier D\ closed\ closure_closed rel_frontier_def subset_eq) let ?w = "(SOME w. w \ rel_interior D)" have w: "?w \ rel_interior D" using \D \ \\ in_rel_interior by blast have "C \ rel_interior D = (D \ C) \ rel_interior D" using rel_interior_subset by blast also have "(D \ C) \ rel_interior D = {}" proof (rule face_of_disjoint_rel_interior) show "D \ C face_of D" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \ \ aff_dim D = int n" using \_def \D \ \\ by blast moreover have "C \ \ \ aff_dim C = int n" using \_def \C \ \\ by blast ultimately show "D \ C \ D" by (metis False face_of_aff_dim_lt inf.idem inf_le1 intface\ not_le poly\ polytope_imp_convex) qed finally have CD: "C \ (rel_interior D) = {}" . have zKC: "(convex hull insert ?z K) \ C" by (metis DiffE \C \ \\ \K \ rel_frontier C\ closed\ closure_closed convex\ hull_minimal insert_subset rel_frontier_def rel_interior_subset subset_iff z) have eq: "convex hull (insert ?z K) \ convex hull (insert ?w L) = convex hull (insert ?z K) \ convex hull L" apply (rule Int_convex_hull_insert_rel_exterior [OF \convex D\ \L \ D\ w]) using zKC CD apply (force simp: disjnt_def) done have ch_id: "convex hull K = K" "convex hull L = L" using "*" \K \ \\ \L \ \\ hull_same by auto have "convex C" by (simp add: \C \ \\ convex\) have "convex hull (insert ?z K) \ L = L \ convex hull (insert ?z K)" by blast also have "... = convex hull K \ L" proof (subst Int_convex_hull_insert_rel_exterior [OF \convex C\ \K \ C\ z]) have "(C \ D) \ rel_interior C = {}" proof (rule face_of_disjoint_rel_interior) show "C \ D face_of C" using \_def \C \ \\ \D \ \\ intface\ by blast have "D \ \" "aff_dim D = int n" using \_def \D \ \\ by fastforce+ moreover have "C \ \" "aff_dim C = int n" using \_def \C \ \\ by fastforce+ ultimately have "aff_dim D + - 1 * aff_dim C \ 0" by fastforce then have "\ C face_of D" using False \convex D\ face_of_aff_dim_lt by fastforce show "C \ D \ C" using \C \ \\ \D \ \\ \\ C face_of D\ intface\ by fastforce qed then have "D \ rel_interior C = {}" by (metis inf.absorb_iff2 inf_assoc inf_sup_aci(1) rel_interior_subset) then show "disjnt L (rel_interior C)" by (meson \L \ D\ disjnt_def disjnt_subset1) next show "L \ convex hull K = convex hull K \ L" by force qed finally have chKL: "convex hull (insert ?z K) \ L = convex hull K \ L" . have "convex hull insert ?z K \ convex hull L face_of K" by (simp add: \K \ \\ \L \ \\ ch_id chKL faceI\) also have "... face_of convex hull insert ?z K" proof - obtain I where I: "\ affine_dependent I" "K = convex hull I" using * [OF \K \ \\] by auto then have "\a. a \ rel_interior C \ a \ affine hull I" using ahK_C_disjoint \C \ \\ \K \ \\ \K \ rel_frontier C\ affine_hull_convex_hull by blast then show ?thesis by (metis I affine_independent_insert face_of_convex_hull_affine_independent hull_insert subset_insertI z) qed finally have 1: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?z K" . have "convex hull insert ?z K \ convex hull L face_of L" by (simp add: \K \ \\ \L \ \\ ch_id chKL faceI\) also have "... face_of convex hull insert ?w L" proof - obtain I where I: "\ affine_dependent I" "L = convex hull I" using * [OF \L \ \\] by auto then have "\a. a \ rel_interior D \ a \ affine hull I" using \D \ \\ \L \ \\ \L \ rel_frontier D\ affine_hull_convex_hull ahK_C_disjoint by blast then show ?thesis by (metis I aff_independent_finite convex_convex_hull face_of_convex_hull_insert face_of_refl hull_insert w) qed finally have 2: "convex hull insert ?z K \ convex hull L face_of convex hull insert ?w L" . show ?thesis by (simp add: X Y eq 1 2) qed qed qed qed then show "S \ S' face_of S" "S \ S' face_of S'" if "S \ \ \ ?\ \ S' \ \ \ ?\" for S S' using that by auto qed show "\F \ \ \ ?\. C = \F" if "C \ \" for C proof (cases "C \ \") case True then show ?thesis by (meson UnCI fin\ subsetD subsetI) next case False then have "C \ \" by (simp add: \_def \_def aff\ less_le that) let ?z = "SOME z. z \ rel_interior C" have z: "?z \ rel_interior C" using \C \ \\ in_rel_interior by blast let ?F = "\K \ \ \ Pow (rel_frontier C). {convex hull (insert ?z K)}" have "?F \ ?\" using \C \ \\ by blast moreover have "C \ \?F" proof fix x assume "x \ C" have "convex C" using \C \ \\ convex\ by blast have "bounded C" using \C \ \\ by (simp add: poly\ polytope_imp_bounded that) have "polytope C" using \C \ \\ poly\ by auto have "\ (?z = x \ C = {?z})" using \C \ \\ aff_dim_sing [of ?z] \\ n \ 1\ by (force simp: \_def) then obtain y where y: "y \ rel_frontier C" and xzy: "x \ closed_segment ?z y" and sub: "open_segment ?z y \ rel_interior C" by (blast intro: segment_to_rel_frontier [OF \convex C\ \bounded C\ z \x \ C\]) then obtain F where "y \ F" "F face_of C" "F \ C" by (auto simp: rel_frontier_of_polyhedron_alt [OF polytope_imp_polyhedron [OF \polytope C\]]) then obtain \ where "finite \" "\ \ \" "F = \\" by (metis (mono_tags, lifting) \_def \C \ \\ \convex C\ aff\ face\ face_of_aff_dim_lt fin\ le_less_trans mem_Collect_eq not_less) then obtain K where "y \ K" "K \ \" using \y \ F\ by blast moreover have x: "x \ convex hull {?z,y}" using segment_convex_hull xzy by auto moreover have "convex hull {?z,y} \ convex hull insert ?z K" by (metis (full_types) \y \ K\ hull_mono empty_subsetI insertCI insert_subset) moreover have "K \ \" using \K \ \\ \\ \ \\ by blast moreover have "K \ rel_frontier C" using \F = \\\ \F \ C\ \F face_of C\ \K \ \\ face_of_subset_rel_frontier by fastforce ultimately show "x \ \?F" by force qed moreover have "convex hull insert (SOME z. z \ rel_interior C) K \ C" if "K \ \" "K \ rel_frontier C" for K proof (rule hull_minimal) show "insert (SOME z. z \ rel_interior C) K \ C" using that \C \ \\ in_rel_interior rel_interior_subset by (force simp: closure_eq rel_frontier_def closed\) show "convex C" by (simp add: \C \ \\ convex\) qed then have "\?F \ C" by auto ultimately show ?thesis by blast qed have "(\C. C \ \ \ L \ C) \ aff_dim L \ int n" if "L \ \ \ ?\" for L using that proof assume "L \ \" then show ?thesis using C\ \_def "*" by fastforce next assume "L \ ?\" then obtain C K where "C \ \" and L: "L = convex hull insert (SOME z. z \ rel_interior C) K" and K: "K \ \" "K \ rel_frontier C" by auto then have "convex hull C = C" by (meson convex\ convex_hull_eq) then have "convex C" by (metis (no_types) convex_convex_hull) have "rel_frontier C \ C" by (metis DiffE closed\ \C \ \\ closure_closed rel_frontier_def subsetI) have "K \ C" using K \rel_frontier C \ C\ by blast have "C \ \" using \_def \C \ \\ by auto moreover have "L \ C" using K L \C \ \\ by (metis \K \ C\ \convex hull C = C\ contra_subsetD hull_mono in_rel_interior insert_subset rel_interior_subset) ultimately show ?thesis using \rel_frontier C \ C\ \L \ C\ aff\ aff_dim_subset \C \ \\ dual_order.trans by blast qed then show "\C. C \ \ \ L \ C" "aff_dim L \ int n" if "L \ \ \ ?\" for L using that by auto qed then show ?thesis apply (rule ex_forward, safe) apply (meson Union_iff subsetCE, fastforce) by (meson infinite_super simplicial_complex_def) qed qed lemma simplicial_subdivision_of_cell_complex_lowdim: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" and aff: "\C. C \ \ \ aff_dim C \ d" obtains \ where "simplicial_complex \" "\K. K \ \ \ aff_dim K \ d" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof (cases "d \ 0") case True then obtain n where n: "d = of_nat n" using zero_le_imp_eq_int by blast have "\\. simplicial_complex \ \ (\K\\. aff_dim K \ int n) \ \\ = \(\C\\. {F. F face_of C}) \ (\C\\C\\. {F. F face_of C}. \F. finite F \ F \ \ \ C = \F) \ (\K\\. \C. C \ (\C\\. {F. F face_of C}) \ K \ C)" proof (rule simplicial_subdivision_aux) show "finite (\C\\. {F. F face_of C})" using \finite \\ poly polyhedron_eq_finite_faces polytope_imp_polyhedron by fastforce show "polytope F" if "F \ (\C\\. {F. F face_of C})" for F using poly that face_of_polytope_polytope by blast show "aff_dim F \ int n" if "F \ (\C\\. {F. F face_of C})" for F using that by clarify (metis n aff_dim_subset aff face_of_imp_subset order_trans) show "F \ (\C\\. {F. F face_of C})" if "G \ (\C\\. {F. F face_of C})" and "F face_of G" for F G using that face_of_trans by blast next show "F1 \ F2 face_of F1 \ F1 \ F2 face_of F2" if "F1 \ (\C\\. {F. F face_of C})" and "F2 \ (\C\\. {F. F face_of C})" for F1 F2 using that by safe (meson face face_of_Int_subface)+ qed moreover have "\(\C\\. {F. F face_of C}) = \\" using face_of_imp_subset face by blast ultimately show ?thesis apply clarify apply (rule that, assumption+) using n apply blast apply (simp_all add: poly face_of_refl polytope_imp_convex) using face_of_imp_subset by fastforce next case False then have m1: "\C. C \ \ \ aff_dim C = -1" by (metis aff aff_dim_empty_eq aff_dim_negative_iff dual_order.trans not_less) then have face\: "\F S. \S \ \; F face_of S\ \ F \ \" by (metis aff_dim_empty face_of_empty) show ?thesis proof have "\S. S \ \ \ \n. n simplex S" by (metis (no_types) m1 aff_dim_empty simplex_minus_1) then show "simplicial_complex \" by (auto simp: simplicial_complex_def \finite \\ face intro: face\) show "aff_dim K \ d" if "K \ \" for K by (simp add: that aff) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C using \C \ \\ equals0I by auto show "\C. C \ \ \ K \ C" if "K \ \" for K using \K \ \\ by blast qed auto qed proposition simplicial_subdivision_of_cell_complex: assumes "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" obtains \ where "simplicial_complex \" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: simplicial_subdivision_of_cell_complex_lowdim [OF assms aff_dim_le_DIM]) corollary fine_simplicial_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" obtains \ where "simplicial_complex \" "\K. K \ \ \ diameter K < e" "\\ = \\" "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" "\K. K \ \ \ \C. C \ \ \ K \ C" proof - obtain \ where \: "finite \" "\\ = \\" and diapoly: "\X. X \ \ \ diameter X < e" "\X. X \ \ \ polytope X" and "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" and \covers: "\C x. C \ \ \ x \ C \ \D. D \ \ \ x \ D \ D \ C" and \covered: "\C. C \ \ \ \D. D \ \ \ C \ D" by (blast intro: cell_complex_subdivision_exists [OF \0 < e\ \finite \\ poly aff_dim_le_DIM face]) then obtain \ where \: "simplicial_complex \" "\\ = \\" and \covers: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and \covered: "\K. K \ \ \ \C. C \ \ \ K \ C" using simplicial_subdivision_of_cell_complex [OF \finite \\] by metis show ?thesis proof show "simplicial_complex \" by (rule \) show "diameter K < e" if "K \ \" for K by (metis le_less_trans diapoly \covered diameter_subset polytope_imp_bounded that) show "\\ = \\" by (simp add: \(2) \\\ = \\\) show "\F. finite F \ F \ \ \ C = \F" if "C \ \" for C proof - { fix x assume "x \ C" then obtain D where "D \ \" "x \ D" "D \ C" using \covers \C \ \\ \covers by force then have "\X\\ \ Pow C. x \ X" using \D \ \\ \D \ C\ \x \ D\ by blast } moreover have "finite (\ \ Pow C)" using \simplicial_complex \\ simplicial_complex_def by auto ultimately show ?thesis by (rule_tac x="(\ \ Pow C)" in exI) auto qed show "\C. C \ \ \ K \ C" if "K \ \" for K by (meson \covered \covered order_trans that) qed qed subsection\Some results on cell division with full-dimensional cells only\ lemma convex_Union_fulldim_cells: assumes "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" and eq: "\\ = U"and "convex U" shows "\{C \ \. aff_dim C = aff_dim U} = U" (is "?lhs = U") proof - have "closed U" using \finite \\ clo eq by blast have "?lhs \ U" using eq by blast moreover have "U \ ?lhs" proof (cases "\C \ \. aff_dim C = aff_dim U") case True then show ?thesis using eq by blast next case False have "closed ?lhs" by (simp add: \finite \\ clo closed_Union) moreover have "U \ closure ?lhs" proof - have "U \ closure(\{U - C |C. C \ \ \ aff_dim C < aff_dim U})" proof (rule Baire [OF \closed U\]) show "countable {U - C |C. C \ \ \ aff_dim C < aff_dim U}" using \finite \\ uncountable_infinite by fastforce have "\C. C \ \ \ openin (top_of_set U) (U-C)" by (metis Sup_upper clo closed_limpt closedin_limpt eq openin_diff openin_subtopology_self) then show "openin (top_of_set U) T \ U \ closure T" if "T \ {U - C |C. C \ \ \ aff_dim C < aff_dim U}" for T using that dense_complement_convex_closed \closed U\ \convex U\ by auto qed also have "... \ closure ?lhs" proof - obtain C where "C \ \" "aff_dim C < aff_dim U" by (metis False Sup_upper aff_dim_subset eq eq_iff not_le) have "\X. X \ \ \ aff_dim X = aff_dim U \ x \ X" if "\V. (\C. V = U - C \ C \ \ \ aff_dim C < aff_dim U) \ x \ V" for x proof - have "x \ U \ x \ \\" using \C \ \\ \aff_dim C < aff_dim U\ eq that by blast then show ?thesis by (metis Diff_iff Sup_upper Union_iff aff_dim_subset dual_order.order_iff_strict eq that) qed then show ?thesis by (auto intro!: closure_mono) qed finally show ?thesis . qed ultimately show ?thesis using closure_subset_eq by blast qed ultimately show ?thesis by blast qed proposition fine_triangular_subdivision_of_cell_complex: assumes "0 < e" "finite \" and poly: "\C. C \ \ \ polytope C" and aff: "\C. C \ \ \ aff_dim C = d" and face: "\C1 C2. \C1 \ \; C2 \ \\ \ C1 \ C2 face_of C1 \ C1 \ C2 face_of C2" obtains \ where "triangulation \" "\k. k \ \ \ diameter k < e" "\k. k \ \ \ aff_dim k = d" "\\ = \\" "\C. C \ \ \ \f. finite f \ f \ \ \ C = \f" "\k. k \ \ \ \C. C \ \ \ k \ C" proof - obtain \ where "simplicial_complex \" and dia\: "\K. K \ \ \ diameter K < e" and "\\ = \\" and in\: "\C. C \ \ \ \F. finite F \ F \ \ \ C = \F" and in\: "\K. K \ \ \ \C. C \ \ \ K \ C" by (blast intro: fine_simplicial_subdivision_of_cell_complex [OF \e > 0\ \finite \\ poly face]) let ?\ = "{K \ \. aff_dim K = d}" show thesis proof show "triangulation ?\" using \simplicial_complex \\ by (auto simp: triangulation_def simplicial_complex_def) show "diameter L < e" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: dia\) show "aff_dim L = d" if "L \ {K \ \. aff_dim K = d}" for L using that by auto show "\F. finite F \ F \ {K \ \. aff_dim K = d} \ C = \F" if "C \ \" for C proof - obtain F where "finite F" "F \ \" "C = \F" using in\ [OF \C \ \\] by auto show ?thesis proof (intro exI conjI) show "finite {K \ F. aff_dim K = d}" by (simp add: \finite F\) show "{K \ F. aff_dim K = d} \ {K \ \. aff_dim K = d}" using \F \ \\ by blast have "d = aff_dim C" by (simp add: aff that) moreover have "\K. K \ F \ closed K \ convex K" using \simplicial_complex \\ \F \ \\ unfolding simplicial_complex_def by (metis subsetCE \F \ \\ closed_simplex convex_simplex) moreover have "convex (\F)" using \C = \F\ poly polytope_imp_convex that by blast ultimately show "C = \{K \ F. aff_dim K = d}" by (simp add: convex_Union_fulldim_cells \C = \F\ \finite F\) qed qed then show "\{K \ \. aff_dim K = d} = \\" by auto (meson in\ subsetCE) show "\C. C \ \ \ L \ C" if "L \ {K \ \. aff_dim K = d}" for L using that by (auto simp: in\) qed qed end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6585 +1,6585 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Line_Segment begin lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally show "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) apply auto done qed (insert \e>0\ \d>0\, auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x]) apply (auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "z \ interior S" apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) by simp (simp add: field_simps norm_minus_commute) then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" apply (clarsimp simp add: min_def a) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a x]) using \0 < e\ False apply (auto simp: field_split_simps) done qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" apply (rule sum.cong) apply rule apply auto done also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "rel_interior (convex hull (insert 0 ?p)) = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 ?p)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 ?p))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 ?p)" using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto then have as [rule_format]: "\y. dist x y < e \ (\i\Basis. i \ D \ y\i = 0) \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e / 2) *\<^sub>R i) \ j" apply - apply (rule as[THEN conjunct1]) using D \e > 0\ x0 apply (auto simp: dist_norm inner_simps inner_Basis) done then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D unfolding dist_norm by auto have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e / 2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e / 2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e / 2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 ?p)" using as assms unfolding substd_simplex[OF assms] by fastforce obtain a where a: "a \ D" using \D \ {}\ by auto let ?d = "(1 - sum ((\) x) D) / real (card D)" have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp) moreover have "?d > 0" using as using \0 < card D\ by auto ultimately have h3: "min (Min (((\) x) ` D)) ?d > 0" by auto have "x \ rel_interior (convex hull (insert 0 ?p))" unfolding rel_interior_ball mem_Collect_eq h0 apply (rule,rule h2) unfolding substd_simplex[OF assms] apply (rule_tac x="min (Min (((\) x) ` D)) ?d" in exI) apply (rule, rule h3) apply safe unfolding mem_ball proof - fix y :: 'a assume y: "dist x y < min (Min ((\) x ` D)) ?d" assume y2: "\i\Basis. i \ D \ y\i = 0" have "sum ((\) y) D \ sum (\i. x\i + ?d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < ?d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] using Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (insert y2, auto) qed } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?D = D let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D" have "finite D" apply (rule finite_subset) using assms(2) apply auto done then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = inverse (2 * real (card D))" apply (rule trans[of _ "sum (\j. if i = j then inverse (2 * real (card D)) else 0) ?D"]) unfolding inner_sum_left apply (rule sum.cong) using \i \ D\ \finite D\ sum.delta'[of D i "(\k. inverse (2 * real (card D)))"] d1 assms(2) by (auto simp: inner_Basis rev_subsetD[OF _ assms(2)]) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real (card D))) ?D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) using \bounded_linear f\ fd * apply auto done ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] apply auto apply blast done moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \x \ a\ apply simp done } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto next case True then have "rel_interior S = {}" by auto then have "closure (rel_interior S) = {}" using closure_empty by auto with True show ?thesis by auto qed qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def apply (simp add: algebra_simps) using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] apply auto done finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" using e_def assms by auto moreover have "e \ 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using 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 subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis apply simp by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) unfolding rel_frontier_def using * closed_affine_hull apply auto done qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" apply (rule affine_dim_equal) using * affine_affine_hull apply auto using \S1 \ {}\ hull_subset[of S1] apply auto using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] apply auto done obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" using x1_def x2_def apply (auto simp add: algebra_simps) using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] apply auto done then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma 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 by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof - { assume "S = {}" then have ?thesis by auto } moreover { assume "T = {}" then have ?thesis by auto } moreover { assume "S \ {}" "T \ {}" then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" - using fst_linear \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto + using linear_fst \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 + using linear_snd \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" apply (subst convex_rel_interior_inter_two[of "S \ (UNIV :: 'm set)" "(UNIV :: 'n set) \ T"]) using * ri assms convex_Times apply auto done finally have ?thesis using * by auto } ultimately show ?thesis by blast qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def 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 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 + using rel_interior_convex_linear_image[of fst S] assms linear_fst 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 + using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] linear_snd conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \ S)" f c x]) using convex_cone_hull[of "{(1 :: real)} \ S"] conv apply auto done { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] apply (subst convex_sum[of I "convex hull \(S ` I)" c s]) using * assms convex_convex_hull apply auto done qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto then have "x \ ?rhs" apply auto apply (rule_tac x = c in exI) apply (rule_tac x = s in exI) using * c_def s_def p \x \ S i\ apply auto done } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def apply auto done moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def apply (rule cInf_greatest [OF nonMT]) using \ dual_order.strict_implies_order le_less_linear by blast with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" apply (rule subsetD [OF rel_interior_subset inint]) using \l \ 0\ \0 < d\ \0 < \\ by auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" apply (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) using 1 2 \0 < d\ \0 < \\ apply (auto simp: algebra_simps) done qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp then show ?thesis apply (rule ray_to_rel_frontier [OF \bounded S\ _ _ \l \ 0\]) using a affine_hull_nonempty_interior apply blast by (simp add: \interior S = rel_interior S\ frontier_def rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy apply (auto simp: in_segment) apply (rule_tac x="d" in exI) using \0 < d\ that apply (auto simp: algebra_simps) done ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" apply (auto simp: in_segment) apply (rule_tac x="1/d" in exI) apply (auto simp: algebra_simps) done next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ x]) using df rel_frontier_def by auto qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i\I. convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" apply (subst sum_set_cond_linear[of convex]) using rel_interior_sum rel_interior_sing[of "0"] assms apply (auto simp add: convex_set_plus) done lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i\I. convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i { fix i assume "i \ I" then have "s i \ S i" using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto } then have "\i\I. s i \ S i" by auto moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" apply (subst convex_hull_finite_union[of I S]) using assms apply blast using assms apply blast apply rule apply (rule_tac x = "(\i. (1 :: real) / of_nat (card I))" in exI) apply auto done } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def apply auto done moreover have "sum A I = S + T" using A_def I_def unfolding set_plus_def apply auto unfolding set_plus_def apply auto done ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) { fix i assume "i \ I" then have "convex (K i)" unfolding K_def apply (subst convex_cone_hull) apply (subst convex_Times) using assms apply auto done } then have convK: "\i\I. convex (K i)" by auto { fix i assume "i \ I" then have "K0 \ K i" unfolding K0_def K_def apply (subst hull_mono) using \\i\I. C0 \ S i\ apply auto done } then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def apply (subst convex_cone_hull) apply (subst convex_Times) unfolding C0_def using convex_convex_hull apply auto done ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" apply (subst cone_convex_hull) using cone_Union[of "K ` I"] apply auto unfolding K_def using cone_cone_hull apply auto done ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" apply (subst convex_hull_finite_union_cones[of I K]) using assms apply blast using False apply blast unfolding K_def apply rule apply (subst convex_cone_hull) apply (subst convex_Times) using assms cone_cone_hull \\i\I. K i \ {}\ K_def apply auto done finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1::real, x) \ rel_interior K0" using K0_def * set_sum_alt[of I "(\i. rel_interior (K i))"] assms k_def cs apply auto apply (rule_tac x = k in exI) apply (simp add: sum_prod) done then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where "0 < e" "ball x e \ interior I" . moreover define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) ultimately have "K \ I" "K < x" "x \ I" using interior_subset[of I] \x \ interior I\ by auto have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "t \ s" shows "convex hull t = affine hull t \ convex hull s" proof - have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto { fix u v x assume uv: "sum u t = 1" "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\v\t. u v *\<^sub>R v)" "x \ t" then have s: "s = (s - t) \ t" \ \split into separate cases\ using assms by auto have [simp]: "(\x\t. v x *\<^sub>R x) + (\x\s - t. v x *\<^sub>R x) = (\x\t. u x *\<^sub>R x)" "sum v t + sum v (s - t) = 1" using uv fin s by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\s. if x \ t then v x - u x else v x) = 0" "(\x\s. (if x \ t then v x - u x else v x) *\<^sub>R x) = 0" using uv fin by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "convex hull t \ affine hull t" using convex_hull_subset_affine_hull by blast moreover have "convex hull t \ convex hull s" using assms hull_mono by blast moreover have "affine hull t \ convex hull s \ convex hull t" using assms apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) apply (drule_tac x=s in spec) apply (auto simp: fin) apply (rule_tac x=u in exI) apply (rename_tac v) apply (drule_tac x="\x. if x \ t then v x - u x else v x" in spec) apply (force)+ done ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "card s = Suc (DIM ('a))" shows "affine hull s = UNIV" proof (cases "s = {}") case True then show ?thesis using assms by simp next case False then obtain a t where t: "a \ t" "s = insert a t" by blast then have fin: "finite t" using assms by (metis finite_insert aff_independent_finite) show ?thesis using assms t fin apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) apply (rule subset_antisym) apply force apply (rule Fun.vimage_subsetD) apply (metis add.commute diff_add_cancel surj_def) apply (rule card_ge_dim_independent) apply (auto simp: card_image inj_on_def dim_subset_UNIV) done qed lemma affine_independent_span_gt: fixes s :: "'a::euclidean_space set" assumes ind: "\ affine_dependent s" and dim: "DIM ('a) < card s" shows "affine hull s = UNIV" apply (rule affine_independent_span_eq [OF ind]) apply (rule antisym) using assms apply auto apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) done lemma empty_interior_affine_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(affine hull s) = {}" using assms apply (induct s rule: finite_induct) apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) apply (rule empty_interior_lowdim) by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) lemma empty_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(convex hull s) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma rel_interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_interior(convex hull s) = {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. s = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ s" "b \ s" "a \ b" then have s: "s = (s - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst s, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull s" "y \ ?rhs" { fix u T a assume ua: "\x\s. 0 \ u x" "sum u s = 1" "\ 0 < u a" "a \ s" and yT: "y = (\x\s. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull s \ {w. \u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = w}" have ua0: "u a = 0" using ua by auto obtain b where b: "b\s" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\s. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\s. u x *\<^sub>R x) \ affine hull s" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull s" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull s" using d e yT by auto then obtain v where "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto then have False using assms apply (simp add: affine_dependent_explicit_finite fs) apply (drule_tac x="\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) using ua b d apply (auto simp: algebra_simps sum_subtractf sum.distrib) done } note * = this have "y \ rel_interior (convex hull s)" using y apply (simp add: mem_rel_interior) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes s :: "'a::euclidean_space set" shows "\ affine_dependent s ==> interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) apply (rule trans [of _ "rel_interior(convex hull s)"]) apply (simp add: affine_independent_span_gt rel_interior_interior) by (simp add: rel_interior_convex_hull_explicit) lemma interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card s" and u: "\x. x\s \ 0 < u x" "sum u s = 1" and a: "a \ s" then have cs: "Suc 0 < card s" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ s" "a \ b" proof (cases "s \ {a}") case True then show thesis using cs subset_singletonD by fastforce next case False then show thesis by (blast intro: that) qed have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u s" apply (rule Groups_Big.sum_mono2) using a b u apply (auto simp: less_imp_le aff_independent_finite assms) done finally have "u a < 1" using \b \ s\ u by fastforce } note [simp] = this show ?thesis using assms apply (auto simp: interior_convex_hull_explicit_minimal) apply (rule_tac x=u in exI) apply (auto simp: not_le) done qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" apply (simp add: segment_convex_hull open_segment_def) apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) done next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" apply simp apply (rule affine_independent_span_gt) apply (simp_all add: affine_dependent_def insert_Diff_if) done then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes s :: "'a::euclidean_space set" shows "compact s ==> closure(convex hull s) = convex hull s" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (\x \ s. u x = 0) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) apply (auto simp: convex_hull_finite fs) apply (drule_tac x=u in spec) apply (rule_tac x=u in exI) apply force apply (rename_tac v) apply (rule notE [OF assms]) apply (simp add: affine_dependent_explicit) apply (rule_tac x=s in exI) apply (auto simp: fs) apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done qed lemma frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (DIM ('a) < card s \ (\x \ s. u x = 0)) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card s") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card s \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = \{convex hull (s - {x}) |x. x \ s}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix u a have "\x\s. 0 \ u x \ a \ s \ u a = 0 \ sum u s = 1 \ \x v. x \ s \ (\x\s - {x}. 0 \ v x) \ sum v (s - {x}) = 1 \ (\x\s - {x}. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ s \ \x\s - {a}. 0 \ u x \ sum u (s - {a}) = 1 \ \v. (\x\s. 0 \ v x) \ (\x\s. v x = 0) \ sum v s = 1 \ (\x\s. v x *\<^sub>R x) = (\x\s - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else rel_frontier(convex hull s))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else \{convex hull (s - {x}) |x. x \ s})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ frontier(convex hull s)" proof (cases "affine_dependent s") case True with assms show ?thesis apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) next case False { assume "card s = Suc (card Basis)" then have cs: "Suc 0 < card s" by (simp) with subset_singletonD have "\y \ s. y \ x" by (cases "s \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ interior(convex hull s)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes s :: "'a::euclidean_space set" assumes "card s = Suc (DIM ('a))" shows "interior(convex hull s) = {} \ affine_dependent s" proof - { fix a b assume ab: "a \ interior (convex hull s)" "b \ s" "b \ affine hull (s - {b})" then have "interior(affine hull s) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then have False using ab by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) } then show ?thesis using assms apply auto apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) apply (auto simp: affine_dependent_def) done qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar s \ \u v w. s \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear s \ (\u v. s \ affine hull {u,v})" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ s" by auto { fix u assume *: "\x y. \x\s; y\s\ \ \c. x - y = c *\<^sub>R u" have "\u v. s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" apply (rule_tac x=x in exI) apply (rule_tac x="x+u" in exI, clarify) apply (erule exE [OF * [OF x]]) apply (rename_tac c) apply (rule_tac x="1+c" in exI) apply (rule_tac x="-c" in exI) apply (simp add: algebra_simps) done } moreover { fix u v x y assume *: "s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "x\s \ y\s \ \c. x - y = c *\<^sub>R (v-u)" apply (drule subsetD [OF *])+ apply simp apply clarify apply (rename_tac r1 r2) apply (rule_tac x="r1-r2" in exI) apply (simp add: algebra_simps) apply (metis scaleR_left.add) done } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD by auto (metis collinear_2 numeral_2_eq_2) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - have "card s \ 2 \ card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis using assms apply safe apply (simp add: collinear_small collinear_imp_coplanar) apply (safe dest!: card_eq_SucD) apply (auto simp: coplanar_def) apply (metis hull_subset insert_subset) done qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar s" "linear f" shows "coplanar(f ` s)" proof - { fix u v w assume "s \ affine hull {u, v, w}" then have "f ` s \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` s \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: "coplanar s \ coplanar ((\x. a + x) ` s)" unfolding coplanar_def apply clarify apply (rule_tac x="u+a" in exI) apply (rule_tac x="v+a" in exI) apply (rule_tac x="w+a" in exI) using affine_hull_translation [of a "{u,v,w}" for u v w] apply (force simp: add.commute) done lemma coplanar_translation_eq: "coplanar((\x. a + x) ` s) \ coplanar s" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" proof assume "coplanar s" then show "coplanar (f ` s)" unfolding coplanar_def using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms by (meson coplanar_def coplanar_linear_image) next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` s)" then obtain u v w where "f ` s \ affine hull {u, v, w}" by (auto simp: coplanar_def) then have "g ` f ` s \ g ` (affine hull {u, v, w})" by blast then have "s \ g ` (affine hull {u, v, w})" using g by (simp add: Fun.image_comp) then show "coplanar s" unfolding coplanar_def using affine_hull_linear_image [of g "{u,v,w}" for u v w] \linear g\ linear_conv_bounded_linear by fastforce qed (*The HOL Light proof is simply MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; *) lemma coplanar_subset: "\coplanar t; s \ t\ \ coplanar s" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: "\collinear {a,b,c}; a \ b\ \ c \ affine hull {a,b}" unfolding collinear_def apply clarify apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (rename_tac y x) apply (simp add: affine_hull_2) apply (rule_tac x="1 - x/y" in exI) apply (simp add: algebra_simps) done lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" apply (case_tac "a=b", simp) apply (case_tac "a=c") apply (simp add: insert_commute) apply (case_tac "b=c") apply (simp add: insert_commute) apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) apply (metis collinear_3_affine_hull insert_commute)+ done lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a,midpoint a b,b}" apply (auto simp: collinear_3 collinear_lemma) apply (drule_tac x="-1" in spec) apply (simp add: algebra_simps) done lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c} " using collinear_midpoint by blast moreover have "collinear{a,b,c} \ b = midpoint a c \ dist a b = dist b c" apply (auto simp: collinear_3_expand assms dist_midpoint) apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) apply (simp add: algebra_simps) done ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False with assms show ?thesis apply (auto simp: collinear_3 collinear_lemma between_norm) apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) apply (simp add: vector_add_divide_simps 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_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" apply (simp add: affine_hull_2, safe) apply (rule_tac x=v in image_eqI) apply (simp add: algebra_simps) apply (metis scaleR_add_left scaleR_one, simp) apply (rule_tac x="1-u" in exI) apply (simp add: algebra_simps) done lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" shows "\\ collinear{a,b,c}; DIM('a) = 2\ \ interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) apply (rule_tac x="u a" in exI, simp) apply (rule_tac x="u b" in exI, simp) apply (rule_tac x="u c" in exI, simp) apply (rename_tac uu x y z) apply (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) apply simp done subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex apply auto[1] using inner_zero_right by blast lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "UNIV \ {x. a \ x = b} \ a = 0 \ b = 0" apply (drule_tac c = "((b+1) / (a \ a)) *\<^sub>R a" in subsetD) apply simp_all by (metis add_cancel_right_right zero_neq_one) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "{x. a \ x < b} \ {} \ a = 0 \ b \ 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "{x. a \ x \ b} \ {} \ a = 0 \ b < 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" by (simp add: open_Collect_less contf) show "open {x. f x < 0}" by (simp add: open_Collect_less contf) show "S \ {x. 0 < f x}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) show "T \ {x. f x < 0}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then show ?thesis apply (rule separation_normal [OF \closed S\]) apply (rule_tac U=U and V=V in that) by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain) using S apply blast by (metis image_iff le_cases nest) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain_gen [of "S k"]) using S apply auto by (meson le_cases nest subsetCE) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case apply (rule_tac x="max m n" in exI, safe) using max.cobounded2 apply blast by (meson le_max_iff_disj) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n apply (rule closed_Int_compact [OF \closed C\]) apply (rule com) using X \K \ S\ \f ` S \ T\ \y \ T\ apply blast apply (rule compact_sequence_with_limit) apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) done then have comf: "compact {a \ K. f a \ insert y (range (\i. f(X(n + i))))}" for n by (simp add: Keq Int_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = {a \ K. f a \ insert y (range (\i. f (X (k + i))))}" apply (rule exE) apply (rule finite_indexed_bound [OF \finite \\ \], assumption, force) done have "X m \ \\" using X le_Suc_ex by (fastforce dest: m) then show ?thesis by blast qed have "\{{a. a \ K \ f a \ insert y (range (\i. f(X(n + i))))} |n. n \ UNIV} \ {}" apply (rule compact_fip_Heine_Borel) using comf apply force using ne apply (simp add: subset_iff del: insert_iff) done then have "\x. x \ (\n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" by blast then show ?thesis apply (simp add: image_iff fX) by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" apply (rule continuous_injective_image_segment_1) apply (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - have "{.. f ` {x,y} \ {}" "{f z<..} \ f ` {x,y} \ {}" using fz fz_notin \x \ S\ \y \ S\ apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm) apply (metis image_eqI less_eq_real_def)+ done then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" apply (rule continuous_injective_image_segment_1) apply (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" apply safe apply (metis (mono_tags, hide_lams) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_inc subsetCE) by (metis (mono_tags, lifting) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis apply (subst *) apply (rule closedin_closed_trans [OF _ closed_UNIV]) apply (rule closedin_compact_projection [OF \compact S\]) by (simp add: clo closedin_closed_Int) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" apply (rule span_mul) apply (rule span_base) apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) apply (auto simp: S T) done with xa image_iff show ?thesis by fastforce qed show "affine hull S \ affine hull (S \ T)" apply (simp add: subset_hull) apply (simp add: \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) apply (force simp: *) done qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp) 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) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True show ?thesis using span_zero [of S] apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) apply (auto simp add: \c = 0\) done next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed show ?thesis apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def cong: image_cong_simp, safe) apply (fastforce simp add: inner_distrib intro: xc_im) apply (force simp: intro!: 2) done qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S apply (simp add: hull_redundant cong: aff_dim_affine_hull2) apply (simp add: affine_hull_insert_span_gen hull_inc) by (force simp add: span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert cong: image_cong_simp) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms show ?thesis apply safe using affine_diffs_subspace [OF assms \b \ S\] apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation image_empty image_insert translation_invert) apply force done qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" apply safe using affine_bounded_eq_trivial assms apply fastforce by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have *: "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" apply (clarsimp simp add: field_split_simps) using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) show ?thesis apply (simp add: C k_def) using ass aa Int_iff empty_iff by blast qed qed have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" apply (rule compact_imp_fip) apply (blast intro: compact_cball) using closed_halfspace_ge k_def apply blast apply (metis *) done then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) show ?thesis apply (rule_tac a=a and b = "a \ z" in that, simp_all) using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen apply blast apply (simp_all add: \a \ 0\ szx) done qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) by (metis \0 < e\ add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) proof - have [simp]: "finite s" "finite t" using aff_independent_finite assms by blast+ have "sum u (s \ t) = 1 \ (\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" if [simp]: "sum u s = 1" "sum v t = 1" and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x have "sum f (s \ t) = 0" apply (simp add: f_def sum_Un sum_subtractf) apply (simp add: sum.inter_restrict [symmetric] Int_commute) done moreover have "(\x\(s \ t). f x *\<^sub>R x) = 0" apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq cong del: if_weak_cong) done ultimately have "\v. v \ s \ t \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ s \ u x = (if x \ t then v x else 0)" by (simp add: f_def) presburger have "sum u (s \ t) = sum u s" by (simp add: sum.inter_restrict) then have "sum u (s \ t) = 1" using that by linarith moreover have "(\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ affine hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: affine_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: convex_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ fixes s :: "'a::euclidean_space set set" assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") proof - have "finite s" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction s rule: finite_induct) case empty then show ?case by auto next case (insert t F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (t \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff that \finite S\ sum_dd0 unfolding affine_dependent_explicit apply (simp only: not_ex) apply (drule_tac x=S in spec) apply (drule_tac x=dd in spec, simp) done consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" apply (rule subset_antisym) using in_convex_hull_exchange_unique assms apply blast by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp) 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 card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis apply (rule_tac f="{}" in that) using True affc apply (simp_all add: eq [symmetric]) by (metis aff_dim_UNIV aff_dim_affine_hull) next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have affeq: "affine hull s = (\x\(\a. c - {a}) ` (c - b). affine hull x)" using \b \ c\ False apply (subst affine_hull_Inter [OF ind, symmetric]) apply (simp add: eq double_diff) done have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed show ?thesis apply (rule_tac f = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" in that) using \finite c\ apply blast apply (simp add: imeq card1 card2) apply (simp add: affeq, clarify) apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) done qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" apply (rule span_mono) using \0 \ S\ add.left_neutral by force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" apply (rule dim_eq_full [THEN iffD1]) apply (rule antisym [OF dim_subset_UNIV]) using DIM_lt apply simp done ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) let ?U = "(+) (c+c) ` {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c + c) ` {x + v |x v. x \ (+) (- c) ` S \ a \ v = 0}" if "u \ S" "b = a \ v" for u v apply (rule_tac x="u+v-c-c" in image_eqI) apply (simp_all add: algebra_simps) apply (rule_tac x="u-c" in exI) apply (rule_tac x="v-c" in exI) apply (simp add: algebra_simps that c) done moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = ?U" by (fastforce simp: algebra_simps) also have "... = range ((+) (c + c))" by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ apply (rule linear_independent_injective_image) using \0 < e\ inj_on_def by fastforce qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" apply (rule dense_complement_openin_affine_hull) apply (simp add: assms rel_interior_aff_dim) using \convex S\ rel_interior_rel_open rel_open by blast then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then have Sclo: "\w. w \ S \ (w + a') \ S" by fastforce show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis apply (rule_tac a' = "a'" and b' = "a' \ z" in that) apply (auto simp: a ba'' inner_left_distrib False Sclo) done qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "{x - a |x. x \ affine hull S} = span {x - a |x. x \ S}" proof - have *: "((\x. x - a) ` (S - {a})) = {x. x + a \ S} - {0}" by (auto simp: algebra_simps) show ?thesis apply (simp add: affine_hull_span2 [OF assms] *) apply (auto simp: algebra_simps) done qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim {x - a |x. x \ S}" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "{x - c |x. x \ S} = {x - a |x. x \ S}" apply safe apply (simp_all only: xy) using mem_affine_3_minus [OF \affine S\] \a \ S\ \c \ S\ apply blast+ done have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim {x - c |x. x \ B}" by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim {x - c | x. x \ affine hull B}" by (simp add: diffs_affine_hull_span \c \ B\) also have "... = dim {x - a |x. x \ S}" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` {x - a| x. x \ T}" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` {x - a| x. x \ T}))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim {x - a| x. x \ T})" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis apply (rule_tac x = T in exI) apply (rule_tac x = "ball x e" in exI) using \T \ \\ apply (simp add: closure_minimal) using closed_cball closure_minimal by blast qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed have 1: "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed have 3: "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast show ?thesis apply (rule_tac x="G (a n)" in exI) apply (intro conjI oG n *) using \K \ S\ \range a = K\ apply blast done qed show ?thesis apply (rule that [OF 1 _ 3]) using Fin \K \ S\ \range a = K\ apply (auto simp: odif) done qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x using D2 [OF that] apply clarify apply (rule_tac x="U \ V" in exI) apply (auto intro: that finite_subset [OF *]) done qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) =(S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis apply (subst eq) apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) by auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" apply (rule closedin_closed_trans) apply (rule continuous_closed_graph_gen [OF contf subset_UNIV]) by (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "open_segment a c \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ open_segment a c" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (open_segment a c))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis using Un_closed_segment [OF b] apply (simp add: closed_segment_eq_open) apply (rule equalityI) using assms apply (simp add: b subset_open_segment) using * by (simp add: insert_commute) qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" apply (intro UN_mono order_refl) apply (simp add: cball_subset_ball_iff field_split_simps) done finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" apply (auto simp: orthogonal_comp_def orthogonal_def) apply (simp add: adjoint_works assms(1) inner_commute) by (metis adjoint_works all_zero_iff assms(1) inner_commute) subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end