diff -r cc1058b83124 -r a5f6d2fc1b1f src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy Thu Jul 27 23:05:25 2023 +0100 +++ b/src/HOL/Analysis/Convex.thy Thu Aug 03 19:10:36 2023 +0200 @@ -22,7 +22,7 @@ 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 + by (simp add: assms convex_def) lemma convexD: assumes "convex s" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" @@ -31,25 +31,7 @@ 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 + by (smt (verit) convexD convexI) lemma convexD_alt: assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" @@ -193,44 +175,31 @@ 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" + and a: "(\ i \ S. a i) = 1" "\i. i \ S \ a i \ 0" + and C: "\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) + using \finite S\ a C +proof (induction 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 + case (insert i S) 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 + case True with insert show ?thesis + by (simp add: sum_nonneg_eq_0_iff) 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\ + using insert + by (simp add: insert.IH flip: sum_divide_distrib) + with \convex C\ insert \0 \ sum a S\ 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) + by (simp add: convex_def) then show ?thesis by (simp add: scaleR_sum_right False) qed @@ -239,18 +208,13 @@ 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 + "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)" + (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + by (metis (full_types) atLeastAtMost_iff convex_sum finite_atLeastAtMost) 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" { @@ -262,18 +226,14 @@ 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 + then have S: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ S" + using sum.If_cases[of "{(1 :: nat) .. 2}" "\x. x = 1" "\x. \" "\x. 1 - \"] + 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] + with sum.atLeast_Suc_atMost have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" - by auto + by (smt (verit, best) Suc_1 Suc_eq_plus1 add_0 le_add1) then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ S" using S by (auto simp: add.commute) } @@ -293,7 +253,7 @@ 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 + by (simp add: convex_sum subsetD) 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" @@ -330,7 +290,8 @@ 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"]] * have "(\x\T. u x *\<^sub>R x) \ S" + with sum[THEN spec[where x="\x. if x\T then u x else 0"]] * + have "(\x\T. u x *\<^sub>R x) \ S" by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) } moreover assume ?rhs ultimately show ?lhs @@ -357,30 +318,14 @@ 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 + by (smt (verit, del_insts) assms mult_cancel_right1 mult_eq_0_iff scaleR_collapse scaleR_eq_0_iff) 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 + by (smt (verit, best) add.commute assms convex_onI distrib_left linorder_cases mult.commute mult_cancel_left2 scaleR_collapse) lemma convex_onD: assumes "convex_on A f" @@ -452,7 +397,7 @@ 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) +proof (clarsimp simp: convex_on_def dist_norm) fix x y assume "x \ S" "y \ S" fix u v :: real @@ -460,20 +405,18 @@ 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 (metis \u + v = 1\ scaleR_left.add scaleR_one) + 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 + then show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" + by (smt (verit, best) \0 \ u\ \0 \ v\ norm_scaleR norm_triangle_ineq) qed subsection\<^marker>\tag unimportant\ \Arithmetic operations on sets preserve convexity\ lemma convex_linear_image: - assumes "linear f" - and "convex S" + assumes "linear f" and "convex S" shows "convex (f ` S)" proof - interpret f: linear f by fact @@ -482,8 +425,7 @@ qed lemma convex_linear_vimage: - assumes "linear f" - and "convex S" + assumes "linear f" and "convex S" shows "convex (f -` S)" proof - interpret f: linear f by fact @@ -494,32 +436,17 @@ 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 + by (simp add: assms convex_linear_image) 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 + by (simp add: assms convex_linear_image) 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 + by (simp add: assms convex_linear_image module_hom_uminus) lemma convex_sums: assumes "convex S" @@ -572,124 +499,79 @@ fixes a :: "'a \ real" and y :: "'a \ 'b::real_vector" and f :: "'b \ real" - assumes "finite s" "s \ {}" + 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))" + 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) +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 \ + case (insert i S) + then have yai: "y i \ C" "a i \ 0" + by auto + with insert 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 + by (simp add: convex_on_def) 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 + with insert have "(\ j \ S. a j) = 0" + by auto + with insert show ?thesis + by (simp add: sum_nonneg_eq_0_iff) 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 ai1: "a i < 1" + using sum_nonneg_leq_bound[of "insert i S" a] insert by force 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 + 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" + have "(\ j \ insert i S. a j) = 1" using insert by auto - then have "(\ j \ s. a j) = 1 - a i" + 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)" + then have "(\ j \ S. a j) / (1 - a i) = 1" 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] + then have a1: "(\ j \ S. ?a j) = 1" + unfolding sum_divide_distrib by simp + 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)" + by (simp add: add.commute insert.hyps) + 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)" + 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))" + also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ S. ?a j *\<^sub>R y j)) + a i * f (y i)" + using ai1 by (smt (verit) asum conv real_scaleR_def yai) + also have "\ \ (1 - a i) * (\ j \ S. ?a j * f (y j)) + a i * f (y i)" + using asum_le i0 by fastforce + also have "\ = (\ j \ S. a j * f (y j)) + a i * f (y i)" + using i0 by (auto simp: sum_distrib_left) + finally show ?thesis using insert by auto - finally show ?thesis - by simp qed qed lemma convex_on_alt: fixes C :: "'a::real_vector set" 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 + (\x \ C. \y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ + f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" + by (smt (verit) convex_on_def) lemma convex_on_diff: fixes f :: "real \ real" @@ -784,12 +666,10 @@ 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" + have "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" + from * have ge: "y - x > 0" "y - x \ 0" and 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\], @@ -798,8 +678,6 @@ 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 @@ -808,75 +686,41 @@ 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 + from z1 have "f x - f y = (x - y) * f' z1" + by (simp add: field_simps) + then have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" + using le(1) z3(3) 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 + with cool' have "f' y - (f x - f y) / (x - y) \ 0" + using z1 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 + by (meson diff_ge_0_iff_ge le(1) neg_divide_le_eq) + have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" + using le(1) z1(3) z2(3) 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" + with z1 assms have "(z1 - x) * f'' z2 \ 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 + using that(3) z1(3) res cool 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 + then show ?thesis + using x y by fastforce 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 0: "\x. x \ C \ f'' x \ 0" + assumes "convex C" + and "\x. x \ C \ DERIV f x :> (f' x)" + and "\x. x \ C \ DERIV f' x :> (f'' x)" + and "\x. x \ C \ f'' x \ 0" shows "convex_on C f" - using f''_imp_f'[OF conv f' f'' 0] assms pos_convex_function - by fastforce + by (metis assms f''_imp_f' pos_convex_function) lemma f''_le0_imp_concave: fixes f :: "real \ real" @@ -965,14 +809,23 @@ qed lemma convex_on_inverse: + fixes A :: "real set" 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) + shows "convex_on A inverse" +proof - + have "convex_on {0::real<..} inverse" + proof (intro convex_on_realI) + fix u v :: real + assume "u \ {0<..}" "v \ {0<..}" "u \ v" + with assms show "-inverse (u^2) \ -inverse (v^2)" + by simp + next + show "\x. x \ {0<..} \ (inverse has_real_derivative - inverse (x\<^sup>2)) (at x)" + by (rule derivative_eq_intros | simp add: power2_eq_square)+ + qed auto + then show ?thesis + using assms convex_on_subset by blast +qed lemma convex_onD_Icc': assumes "convex_on {x..y} f" "c \ {x..y}" @@ -995,7 +848,7 @@ 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) +qed (use assms in auto) lemma convex_onD_Icc'': assumes "convex_on {x..y} f" "c \ {x..y}" @@ -1018,7 +871,7 @@ 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) +qed (use assms in auto) subsection \Some inequalities\ @@ -1161,59 +1014,34 @@ lemma cone_iff: assumes "S \ {}" - shows "cone S \ 0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" -proof - + shows "cone S \ 0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" (is "_ = ?rhs") +proof + assume "cone S" { - 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 \0 < c\ \cone S\ mem_cone by fastforce - } - 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 + fix c :: real + assume "c > 0" + have "x \ ((*\<^sub>R) c) ` S" if "x \ S" for x + using \cone S\ \c>0\ mem_cone[of S x "1/c"] that + exI[of "(\t. t \ S \ x = c *\<^sub>R t)" "(1 / c) *\<^sub>R x"] + by auto + then have "((*\<^sub>R) c) ` S = S" + using \0 < c\ \cone S\ mem_cone by fastforce } - 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 + then show "0 \ S \ (\c. c > 0 \ ((*\<^sub>R) c) ` S = S)" + using \cone S\ cone_contains_0[of S] assms by auto +next + show "?rhs \ cone S" + by (metis Convex.cone_def imageI order_neq_le_trans scaleR_zero_left) 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)) + by (metis cone_hull_empty hull_subset subset_empty) 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 + by (metis cone_cone_hull cone_contains_0 cone_hull_empty_iff) lemma mem_cone_hull: assumes "x \ S" "c \ 0" @@ -1222,65 +1050,27 @@ 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" - using zero_le_one by blast - 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 +proof + have "?rhs \ Collect cone" + using Convex.cone_def by fastforce + moreover have "S \ ?rhs" + by (smt (verit) mem_Collect_eq scaleR_one subsetI) + ultimately show "?lhs \ ?rhs" + using hull_minimal by blast +qed (use mem_cone_hull in auto) lemma convex_cone: - "convex s \ cone s \ (\x\s. \y\s. (x + y) \ s) \ (\x\s. \c\0. (c *\<^sub>R x) \ s)" + "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" + assume "x\S" "y\S" and ?lhs + then have "2 *\<^sub>R x \S" "2 *\<^sub>R y \ S" "convex 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 have "x + y \ S" + using convexD [OF \convex S\, of "2*\<^sub>R x" "2*\<^sub>R y"] + by (smt (verit, ccfv_threshold) field_sum_of_halves scaleR_2 scaleR_half_double) } then show ?thesis unfolding convex_def cone_def by blast @@ -1315,13 +1105,12 @@ qed corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)" -by (simp add: convex_connected) + 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) + using assms by (auto simp: inner_add_left convex_def) lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (\i\Basis. 0 \ x\i)}" by (rule convex_prod) (simp flip: atLeast_def) @@ -1329,9 +1118,7 @@ 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 + by (metis (mono_tags) convex_Inter hull_def mem_Collect_eq) lemma convex_hull_subset: "s \ convex hull t \ convex hull s \ convex hull t" @@ -1344,56 +1131,50 @@ lemma convex_hull_linear_image: assumes f: "linear f" - shows "f ` (convex hull s) = convex hull (f ` s)" + shows "f ` (convex hull S) = convex hull (f ` S)" proof - show "convex hull (f ` s) \ f ` (convex hull s)" + 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 + show "f ` (convex hull S) \ convex hull (f ` S)" + by (meson convex_convex_hull convex_linear_vimage f hull_minimal hull_subset image_subset_iff_subset_vimage) 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 + assumes "linear f" "x \ convex hull S" + shows "f x \ convex hull (f ` S)" + using assms convex_hull_linear_image image_eqI by blast lemma convex_hull_Times: - "convex hull (s \ t) = (convex hull s) \ (convex hull t)" + "convex hull (S \ T) = (convex hull S) \ (convex hull T)" proof - show "convex hull (s \ t) \ (convex hull s) \ (convex hull t)" + 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 + 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)" + 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))" + 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)}" + 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)}" . + finally show "convex {y. (x, y) \ convex hull (S \ T)}" . next - show "convex {x. (x, y) \ convex hull s \ t}" + 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))" + 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)}" + 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)}" . + finally show "convex {x. (x, y) \ convex hull (S \ T)}" . qed qed - then show "(convex hull s) \ (convex hull t) \ convex hull (s \ t)" + then show "(convex hull S) \ (convex hull T) \ convex hull (S \ T)" unfolding subset_eq split_paired_Ball_Sigma by blast qed @@ -1401,10 +1182,10 @@ subsubsection\<^marker>\tag unimportant\ \Stepping theorems for convex hulls of finite sets\ lemma convex_hull_empty[simp]: "convex hull {} = {}" - by (rule hull_unique) auto + by (simp add: hull_same) lemma convex_hull_singleton[simp]: "convex hull {a} = {a}" - by (rule hull_unique) auto + by (simp add: hull_same) lemma convex_hull_insert: fixes S :: "'a::real_vector set" @@ -1415,19 +1196,17 @@ 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)" + then show "x \ ?hull" unfolding insert_iff proof assume "x = a" then show ?thesis - by (rule_tac x=1 in exI) (use assms hull_subset in fastforce) + by (smt (verit, del_insts) add.right_neutral assms ex_in_conv hull_inc mem_Collect_eq scaleR_one scaleR_zero_left) next assume "x \ S" - with hull_subset[of S convex] show ?thesis + with hull_subset show ?thesis by force qed - then show "x \ ?hull" - by simp next fix x assume "x \ ?hull" @@ -1468,44 +1247,35 @@ 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) + by (simp add: as(3)) also have "\ = u * v1 + v * v2" - by simp - finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" by auto + by (smt (verit, ccfv_SIG) distrib_left mult_cancel_left1 obt1(3) obt2(3)) + finally have **:"1 - (u * u1 + v * u2) = u * v1 + v * v2" . 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 + using as obt1 obt2 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) + using False mem_convex_alt obt1(4) obt2(4) zeroes(2) zeroes(4) by fastforce 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 + obtain u1: "u1 \ 1" and u2: "u2 \ 1" + using obt1 obt2 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 + by (smt (verit, ccfv_SIG) as mult_right_mono) 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)) + by (simp add: as 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\) @@ -1516,9 +1286,9 @@ "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 - using diff_add_cancel diff_ge_0_iff_ge by blast + apply (simp add: convex_hull_insert) + using diff_add_cancel diff_ge_0_iff_ge + by (smt (verit, del_insts) Collect_cong) subsubsection\<^marker>\tag unimportant\ \Explicit expression for convex hull\ @@ -1608,87 +1378,56 @@ 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 - +proof (intro subset_antisym subsetI) + fix x + assume "x \ convex hull p" + 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 { - 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) - by (metis (no_types, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg) - } - 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 + fix j + assume "j\{1..k}" + then have "y j \ p \ 0 \ sum u {i. Suc 0 \ i \ i \ k \ y i = y j}" + by (metis (mono_tags, lifting) One_nat_def atLeastAtMost_iff mem_Collect_eq obt(1) sum_nonneg) } - moreover + 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" + by (smt (verit, ccfv_SIG) imageE mem_Collect_eq obt(1) subsetI sum.cong sum.infinite sum_nonneg) + then show "x \ ?rhs" by auto +next + fix y + assume "y \ ?rhs" + then obtain S u where + S: "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 S(1)] unfolding bij_betw_def by auto + then have "0 \ u (f i)" "f i \ p" if "i \ {1..card S}" for i + using S \i \ {1..card S}\ by blast+ + 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}" - using f(1) inj_onD by fastforce - 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 + assume "y\S" + then obtain i where "i\{1..card S}" "f i = y" + by (metis f(2) image_iff) + then have "{x. Suc 0 \ x \ x \ card S \ f x = y} = {i}" + using f(1) inj_onD by fastforce + 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 (simp_all add: sum_constant_scaleR \f i = y\) } - ultimately show ?thesis - unfolding set_eq_iff by blast + then have "(\x = 1..card S. u (f x)) = 1" "(\i = 1..card S. u (f i) *\<^sub>R f i) = y" + by (metis (mono_tags, lifting) S(4,5) f sum.reindex_cong)+ + ultimately + show "y \ convex hull p" + unfolding convex_hull_indexed + by (smt (verit, del_insts) mem_Collect_eq sum.cong) qed @@ -1832,10 +1571,7 @@ 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 + by (smt (verit) aff_dim_affine_hull aff_dim_subset convex_hull_subset_affine_hull hull_subset) subsection \Caratheodory's theorem\ @@ -1856,40 +1592,31 @@ by (rule_tac ex_least_nat_le, auto) 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 + then obtain S u where S: "finite S" "card S = n" "S\p" + and u: "\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 + by (smt (verit) independent_card_le_aff_dim S(3)) 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 + using affine_dependent_explicit_finite[OF S(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" - by (meson Diff_iff sum_nonneg) - then have "sum w S > 0" - using as obt(1) sum_nonneg_eq_0_iff wv by blast - then show False using wv(1) by auto - qed + by (smt (verit, best) S(1) sum_pos2 wv) then have "i \ {}" unfolding i_def by auto then have "t \ 0" - using Min_ge_iff[of i 0] and obt(1) + using Min_ge_iff[of i 0] and S(1) u[unfolded le_less] 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 + using u(1) by blast show "0 \ u v + t * w v" proof (cases "w v < 0") case False @@ -1897,26 +1624,26 @@ next case True then have "t \ u v / (- w v)" - using \v\S\ obt unfolding t_def i_def by (auto intro: Min_le) + using \v\S\ S unfolding t_def i_def by (auto intro: Min_le) then show ?thesis unfolding real_0_le_add_iff using True neg_le_minus_divide_eq 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 + using Min_in[OF _ \i\{}\] and S(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 + unfolding sum.remove[OF S(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 + by (metis add.right_neutral mult_zero_right sum.distrib sum_distrib_left u(2) wv(1)) 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) + unfolding sum.distrib u(3) 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 + using S t a apply (auto simp: * scaleR_left_distrib) done then show False @@ -1924,7 +1651,7 @@ 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 + using S u by auto qed auto lemma caratheodory_aff_dim: @@ -1934,7 +1661,7 @@ proof have "\x S u. \finite S; S \ p; int (card S) \ aff_dim p + 1; \x\S. 0 \ u x; sum u S = 1\ \ (\v\S. u v *\<^sub>R v) \ convex hull S" - by (simp add: hull_subset convex_explicit [THEN iffD1, OF convex_convex_hull]) + by (metis (mono_tags, lifting) convex_convex_hull convex_explicit hull_subset) then show "?lhs \ ?rhs" by (subst convex_hull_caratheodory_aff_dim, auto) qed (use hull_mono in auto) @@ -1986,23 +1713,19 @@ 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 + by (simp add: set_plus_image linear_iff scaleR_right_distrib convex_hull_Times + flip: convex_hull_linear_image) 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) + by (simp add: convex_hull_set_plus translation_eq_singleton_plus) 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]) + by (simp add: convex_hull_linear_image) lemma convex_hull_affinity: "convex hull ((\x. a + c *\<^sub>R x) ` S) = (\x. a + c *\<^sub>R x) ` (convex hull S)" @@ -2022,31 +1745,25 @@ 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" + by (simp_all add: cone_cone_hull mem_cone uv xy) + then obtain cx :: real and xx + and cy :: real and yy where x: "u *\<^sub>R x = cx *\<^sub>R xx" "cx \ 0" "xx \ S" + and 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 - } + + have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" if "cx + cy \ 0" + using "*"(1) nless_le that x(2) y by fastforce 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 + have "u *\<^sub>R x + v *\<^sub>R y \ cone hull S" if "cx + cy > 0" + proof - + 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 that 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" + then show ?thesis using x y by auto - } + qed 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 @@ -2054,28 +1771,7 @@ 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 + by (metis (no_types, lifting) affine_hull_convex_hull affine_hull_eq_empty assms cone_iff convex_hull_scaling hull_inc) subsection \Radon's theorem\ @@ -2084,29 +1780,17 @@ 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 - by (auto simp: Int_absorb1 sum.inter_restrict[OF \finite c\, symmetric]) -qed + using affine_dependent_explicit_finite assms by blast 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" + 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 + then show ?thesis + using assms by (simp add: sum.inter_filter flip: sum.distrib add_eq_0_iff) qed lemma Radon_v_lemma: @@ -2115,19 +1799,15 @@ 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 + have "\x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x" + using assms by auto + then show ?thesis + using assms by (simp add: sum.inter_filter eq_neg_iff_add_eq_0 flip: sum.distrib add_eq_0_iff) qed lemma Radon_partition: assumes "finite C" "affine_dependent C" - shows "\m p. m \ p = {} \ m \ p = C \ (convex hull m) \ (convex hull p) \ {}" + 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 @@ -2139,19 +1819,12 @@ 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" - by (rule_tac sum_mono, auto) - 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) - + by (smt (verit) assms(1) fin(1) mem_Collect_eq sum.neutral_const sum_mono_inv uv) + next + case True + with fin uv show "sum u {x \ C. 0 < u x} \ 0" + by (smt (verit) fin(1) mem_Collect_eq sum_nonneg_eq_0_iff uv) + qed then have *: "sum u {x\C. u x > 0} > 0" unfolding less_le by (metis (no_types, lifting) mem_Collect_eq sum_nonneg) moreover have "sum u ({x \ C. 0 < u x} \ {x \ C. u x < 0}) = sum u C" @@ -2188,97 +1861,85 @@ 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 - with S show ?thesis - by (force intro: that[of p m]) -qed + obtains M P where "M \ c" "P \ c" "M \ P = {}" "(convex hull M) \ (convex hull P) \ {}" + by (smt (verit) Radon_partition affine_dependent_explicit affine_dependent_explicit_finite assms le_sup_iff) subsection \Helly's theorem\ lemma Helly_induct: - fixes f :: "'a::euclidean_space set set" - assumes "card f = n" + fixes \ :: "'a::euclidean_space set set" + assumes "card \ = n" and "n \ DIM('a) + 1" - and "\s\f. convex s" "\t\f. card t = DIM('a) + 1 \ \t \ {}" - shows "\f \ {}" + and "\S\\. convex S" "\T\\. card T = DIM('a) + 1 \ \T \ {}" + shows "\\ \ {}" using assms -proof (induction n arbitrary: f) +proof (induction n arbitrary: \) 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 \ {}" + have "finite \" + using \card \ = Suc n\ by (auto intro: card_ge_0_finite) + show "\\ \ {}" proof (cases "n = DIM('a)") case True then show ?thesis - by (simp add: Suc.prems(1) Suc.prems(4)) + by (simp add: Suc.prems) next case False - have "\(f - {s}) \ {}" if "s \ f" for s + have "\(\ - {S}) \ {}" if "S \ \" for S proof (rule Suc.IH[rule_format]) - show "card (f - {s}) = n" - by (simp add: Suc.prems(1) \finite f\ that) + show "card (\ - {S}) = n" + by (simp add: Suc.prems(1) \finite \\ that) show "DIM('a) + 1 \ n" using False Suc.prems(2) by linarith - show "\t. \t \ f - {s}; card t = DIM('a) + 1\ \ \t \ {}" + show "\t. \t \ \ - {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})" + then have "\S\\. \x. x \ \(\ - {S})" by blast - then obtain X where X: "\s. s\f \ X s \ \(f - {s})" + then obtain X where X: "\S. S\\ \ X S \ \(\ - {S})" by metis show ?thesis - proof (cases "inj_on X f") + proof (cases "inj_on X \") case False - then obtain s t where "s\t" and st: "s\f" "t\f" "X s = X t" + then obtain S T where "S\T" and st: "S\\" "T\\" "X S = X T" unfolding inj_on_def by auto - then have *: "\f = \(f - {s}) \ \(f - {t})" by auto + then have *: "\\ = \(\ - {S}) \ \(\ - {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 + then obtain M P where mp: "M \ P = {}" "M \ P = X ` \" "convex hull M \ convex hull P \ {}" + using Radon_partition[of "X ` \"] and affine_dependent_biggerset[of "X ` \"] + unfolding card_image[OF True] and \card \ = Suc n\ + using Suc(3) \finite \\ and False by auto - have "m \ X ` f" "p \ X ` f" + have "M \ X ` \" "P \ X ` \" using mp(2) by auto - then obtain g h where gh:"m = X ` g" "p = X ` h" "g \ f" "h \ f" + then obtain \ \ where gh:"M = X ` \" "P = X ` \" "\ \ \" "\ \ \" 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 + then have "\ \ (\ \ \) = \" by auto + then have \: "\ = \ \ \" + using inj_on_Un_image_eq_iff[of X \ "\ \ \"] and True unfolding mp(2)[unfolded image_Un[symmetric] gh] by auto - have *: "g \ h = {}" - using gh(1) gh(2) local.mp(1) by blast - 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\)+ + have *: "\ \ \ = {}" + using gh local.mp(1) by blast + have "convex hull (X ` \) \ \\" "convex hull (X ` \) \ \\" + by (rule hull_minimal; use X * \ in \auto simp: Suc.prems(3) convex_Inter\)+ then show ?thesis - unfolding f using mp(3)[unfolded gh] by blast + unfolding \ 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 \ {}" + fixes \ :: "'a::euclidean_space set set" + assumes "card \ \ DIM('a) + 1" "\s\\. convex s" + and "\t. \t\\; card t = DIM('a) + 1\ \ \t \ {}" + shows "\\ \ {}" using Helly_induct assms by blast subsection \Epigraphs of convex functions\ @@ -2371,25 +2032,18 @@ 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 + by (metis assms convex_hull_eq convex_hull_set_plus) 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 + using assms + by (induction A rule: infinite_finite_induct) (auto simp: convex_set_plus) 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) + assumes "\i\A. finite (B i)" shows "finite (\i\A. B i)" + using assms + by (induction A rule: infinite_finite_induct) (auto simp: finite_set_plus) lemma box_eq_set_sum_Basis: "{x. \i\Basis. x\i \ B i} = (\i\Basis. (\x. x *\<^sub>R i) ` (B i))" (is "?lhs = ?rhs") @@ -2401,7 +2055,7 @@ have "sum f Basis \ i \ B i" if "i \ Basis" and f: "\i\Basis. f i \ (\x. x *\<^sub>R i) ` B i" for i f proof - have "(\x\Basis - {i}. f x \ i) = 0" - proof (rule sum.neutral, intro strip) + proof (intro strip sum.neutral) show "f x \ i = 0" if "x \ Basis - {i}" for x using that f \i \ Basis\ inner_Basis that by fastforce qed @@ -2418,10 +2072,6 @@ 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 - + by (induction A rule: infinite_finite_induct) (auto simp: convex_hull_set_plus) end \ No newline at end of file