# HG changeset patch # User paulson # Date 1602012869 -3600 # Node ID 4a2c0eb482aaac65481a87fc2f4cd488b0bde1a3 # Parent 698b58513fd1021d9ad5c6549183b3313cec57c8 Simplified some proofs diff -r 698b58513fd1 -r 4a2c0eb482aa src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy Mon Oct 05 22:53:40 2020 +0100 +++ b/src/HOL/Analysis/Convex.thy Tue Oct 06 20:34:29 2020 +0100 @@ -59,10 +59,8 @@ 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 + by (simp add: convex_def zero_le_divide_iff add_divide_distrib [symmetric]) lemma convex_empty[intro,simp]: "convex {}" unfolding convex_def by simp @@ -193,72 +191,72 @@ lemma convex_sum: fixes C :: "'a::real_vector set" - assumes "finite s" + 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" + 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" + 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 "\j\S. 0 \ a j" and "y i \ C" - and "\j\s. y j \ C" + and "\j\S. y j \ C" using insert.hyps(1,2) insert.prems by simp_all - then have "0 \ sum a 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") + 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" + 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" + 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\ + 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" + 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\ + 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" + 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\ + 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)" + "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" + 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" + 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" + 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 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" @@ -269,43 +267,43 @@ 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" + 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 have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ S" + using S by (auto simp: add.commute) } - then show "convex s" + 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)" + 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" + 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 + 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" + 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" + 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 @@ -321,30 +319,30 @@ 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)" + 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)" + (is "?lhs = ?rhs") +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 + fix T :: "'a set" and u :: "'a \ real" + 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"]] * 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 + unfolding convex_explicit by auto +qed (auto simp: convex_explicit assms) 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)" + 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 \ @@ -388,17 +386,17 @@ 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" +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)" + 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" + assume "x \ S" "y \ S" moreover fix u v :: real assume "0 \ u" "0 \ v" "u + v = 1" @@ -415,8 +413,8 @@ lemma convex_on_cmul [intro]: fixes c :: real assumes "0 \ c" - and "convex_on s f" - shows "convex_on s (\x. c * f x)" + 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 @@ -426,9 +424,9 @@ qed lemma convex_lower: - assumes "convex_on s f" - and "x \ s" - and "y \ s" + assumes "convex_on S f" + and "x \ S" + and "y \ S" and "0 \ u" and "0 \ v" and "u + v = 1" @@ -444,11 +442,11 @@ qed lemma convex_on_dist [intro]: - fixes s :: "'a::real_normed_vector set" - shows "convex_on s (\x. dist a x)" + 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" + assume "x \ S" "y \ S" fix u v :: real assume "0 \ u" assume "0 \ v" @@ -467,32 +465,32 @@ lemma convex_linear_image: assumes "linear f" - and "convex s" - shows "convex (f ` s)" + and "convex S" + shows "convex (f ` S)" proof - interpret f: linear f by fact - from \convex s\ show "convex (f ` s)" + 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)" + and "convex S" + shows "convex (f -` S)" proof - interpret f: linear f by fact - from \convex s\ show "convex (f -` s)" + 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)" + 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) + using \convex S\ by (rule convex_linear_image) qed lemma convex_scaled: @@ -656,7 +654,6 @@ 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)" @@ -720,7 +717,7 @@ 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)] + unfolding convex_on_alt using assms proof safe fix x y \ :: real @@ -737,7 +734,7 @@ 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 + by auto qed lemma atMostAtLeast_subset_convex: @@ -1081,15 +1078,11 @@ subsubsection \Conic hull\ -lemma cone_cone_hull: "cone (cone hull s)" +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 cone_hull_eq: "cone hull S = S \ cone S" + by (metis cone_cone_hull hull_same) lemma mem_cone: assumes "cone S" "x \ S" "c \ 0" @@ -1099,15 +1092,7 @@ 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 + using assms mem_cone by fastforce lemma cone_0: "cone {0}" unfolding cone_def by auto @@ -1138,8 +1123,7 @@ 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 + using \0 < c\ \cone S\ mem_cone by fastforce } ultimately have "((*\<^sub>R) c) ` S = S" by blast } @@ -1202,9 +1186,7 @@ fix x assume "x \ S" then have "1 *\<^sub>R x \ ?rhs" - apply auto - apply (rule_tac x = 1 in exI, auto) - done + using zero_le_one by blast then have "x \ ?rhs" by auto } then have "S \ ?rhs" by auto @@ -1477,7 +1459,7 @@ 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) + using diff_add_cancel diff_ge_0_iff_ge by blast subsubsection\<^marker>\tag unimportant\ \Explicit expression for convex hull\ @@ -1580,13 +1562,9 @@ { fix j assume "j\{1..k}" - then have "y j \ p" "0 \ sum u {i. Suc 0 \ i \ i \ k \ y i = y j}" + 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 + 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" @@ -1611,7 +1589,6 @@ 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}" @@ -1627,9 +1604,7 @@ 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) + 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" @@ -1667,71 +1642,65 @@ "(\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" +proof (cases "a \ S") + case True 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 + show ?thesis + proof + assume ?lhs + then show ?rhs + unfolding * by force + next + 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 + then show ?lhs + using uv True assms + apply (rule_tac x = "\x. (if a = x then v else 0) + u x" in exI) + apply (auto simp: sum_clauses scaleR_left_distrib sum.distrib sum_delta''[OF fin]) + done + qed 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]) + case False + show ?thesis + proof + 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 + then show ?rhs + using u \a\S\ by (rule_tac x="u a" in exI) (auto simp: sum_clauses assms) + 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 + 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 False by (auto intro!: sum.cong) + ultimately show ?lhs + using False by (rule_tac x="\x. if a = x then v else u x" in exI) (auto simp: sum_clauses(2)[OF assms]) + qed 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}" +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}" + (is "?lhs = ?rhs") 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) + have "\x v u. \0 \ v; v \ 1; (1 - v) *\<^sub>R b = x - v *\<^sub>R a\ + \ \u v. x = u *\<^sub>R a + v *\<^sub>R b \ 0 \ u \ 0 \ v \ u + v = 1" + by (metis add.commute diff_add_cancel diff_ge_0_iff_ge) + moreover + have "\u v. \0 \ u; 0 \ v; u + v = 1\ + \ \p\0. \q. 0 \ q b \ q b = 1 - p \ q b *\<^sub>R b = u *\<^sub>R a + v *\<^sub>R b - p *\<^sub>R a" apply (rule_tac x=u in exI, simp) apply (rule_tac x="\x. v" in exI, simp) done + ultimately show ?thesis + using convex_hull_finite_step[OF **, of a 1] + by (auto simp add: convex_hull_finite) qed lemma convex_hull_2_alt: "convex hull {a,b} = {a + u *\<^sub>R (b - a) | u. 0 \ u \ u \ 1}" @@ -1742,11 +1711,8 @@ 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 + apply (simp add: *) + by (rule ex_cong1) (auto simp: algebra_simps) qed lemma convex_hull_3: @@ -1818,54 +1784,51 @@ 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}" + {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" + 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 + 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 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" + 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" + 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" + 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 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" + 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 + 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 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 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" + have t: "\v\S. u v + t * w v \ 0" proof fix v - assume "v \ s" + 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" @@ -1875,28 +1838,24 @@ 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 + using \v\S\ obt unfolding t_def i_def by (auto intro: Min_le) 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 + 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" + 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" + 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" + 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="(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) @@ -1904,69 +1863,48 @@ 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" + 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}" + 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 + 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]) + then show "?lhs \ ?rhs" + by (subst convex_hull_caratheodory_aff_dim, auto) +qed (use hull_mono in auto) 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}" + {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 + unfolding convex_hull_caratheodory_aff_dim + using aff_dim_le_DIM [of p] by fastforce +qed (auto simp: convex_hull_explicit) theorem caratheodory: "convex hull p = - {x::'a::euclidean_space. \s. finite s \ s \ p \ - card s \ DIM('a) + 1 \ x \ convex hull s}" + {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" + 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 + then show "\S. finite S \ S \ p \ card S \ DIM('a) + 1 \ x \ convex hull S" + using convex_hull_finite by fastforce +qed (use hull_mono in force) subsection\<^marker>\tag unimportant\\Some Properties of subset of standard basis\ @@ -1989,7 +1927,7 @@ lemma convex_hull_set_plus: "convex hull (S + T) = convex hull S + convex hull T" - unfolding set_plus_image + 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) @@ -2009,7 +1947,7 @@ 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) + by (metis convex_hull_scaling convex_hull_translation image_image) subsection\<^marker>\tag unimportant\ \Convexity of cone hulls\ @@ -2089,20 +2027,19 @@ 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" + 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 + 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 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}" + 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 @@ -2114,10 +2051,10 @@ qed lemma Radon_v_lemma: - assumes "finite s" - and "sum f s = 0" + 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}" + 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 @@ -2130,79 +2067,63 @@ qed lemma Radon_partition: - assumes "finite c" "affine_dependent c" - shows "\m p. m \ p = {} \ m \ p = c \ (convex hull m) \ (convex hull p) \ {}" + 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" + 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}" + 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" + 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") + 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 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) - 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)" + 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" + "(\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)" + by (rule_tac[!] sum.mono_neutral_left, auto) + 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}" + moreover have "\x\{v \ C. u v < 0}. 0 \ inverse (sum u {x \ C. 0 < u x}) * - u x" + using * by (fastforce intro: mult_nonneg_nonneg) + 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}" + 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] + by (auto simp: z_def sum_negf sum_distrib_left[symmetric]) + moreover have "\x\{v \ C. 0 < u v}. 0 \ inverse (sum u {x \ C. 0 < u x}) * u x" + using * by (fastforce intro: mult_nonneg_nonneg) + 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) + 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 + unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] + using * by (auto simp: z_def sum_negf sum_distrib_left[symmetric]) 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) + 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 @@ -2211,19 +2132,16 @@ 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" + 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" + 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 \ {}" + 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 + with S show ?thesis + by (force intro: that[of p m]) qed @@ -2288,10 +2206,7 @@ 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 + 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\)+ then show ?thesis @@ -2305,10 +2220,7 @@ 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 + using Helly_induct assms by blast subsection \Epigraphs of convex functions\ @@ -2323,14 +2235,7 @@ 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 + using L by (fastforce simp: convex_def convex_on_def epigraph_def) next assume "convex_on S f" "convex S" then show "convex (epigraph S f)" @@ -2350,53 +2255,59 @@ 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})" + (is "?lhs = (\k u x. ?rhs k u x)") +proof + assume ?lhs + then have \
: "convex {xy. fst xy \ S \ f (fst xy) \ snd xy}" + by (metis assms convex_epigraph epigraph_def) + show "\k u x. ?rhs k u x" + proof (intro allI) + fix k u x + show "?rhs k u x" + using \
+ unfolding convex mem_Collect_eq fst_sum snd_sum + 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 + done + qed +next + assume "\k u x. ?rhs k u x" + then show ?lhs + unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq fst_sum snd_sum + using assms[unfolded convex] apply clarsimp + apply (rule_tac y="\i = 1..k. u i * f (fst (x i))" in order_trans) + by (auto simp add: mult_left_mono intro: sum_mono) +qed - 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" + 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" + 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" + u: "\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 + unfolding sum_distrib_right[symmetric] u(2) mult_1 + using assms(2) mult_left_mono u(1) by blast 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 + using hull_inc u by fastforce qed lemma inner_sum_Basis[simp]: "i \ Basis \ (\Basis) \ i = 1" @@ -2425,22 +2336,29 @@ 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 + "{x. \i\Basis. x\i \ B i} = (\i\Basis. (\x. x *\<^sub>R i) ` (B i))" (is "?lhs = ?rhs") +proof - + have "\x. \i\Basis. x \ i \ B i \ + \s. x = sum s Basis \ (\i\Basis. s i \ (\x. x *\<^sub>R i) ` B i)" + by (metis (mono_tags, lifting) euclidean_representation image_iff) + moreover + 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) + show "f x \ i = 0" if "x \ Basis - {i}" for x + using that f \i \ Basis\ inner_Basis that by fastforce + qed + then have "(\x\Basis. f x \ i) = f i \ i" + by (metis (no_types) \i \ Basis\ add.right_neutral sum.remove [OF finite_Basis]) + then have "(\x\Basis. f x \ i) \ B i" + using f that(1) by auto + then show ?thesis + by (simp add: inner_sum_left) + qed + ultimately show ?thesis + by (subst set_sum_alt [OF finite_Basis]) auto +qed lemma convex_hull_set_sum: "convex hull (\i\A. B i) = (\i\A. convex hull (B i))"