# HG changeset patch # User wenzelm # Date 1433968136 -7200 # Node ID 5035a2af185b613a9bd5e6084fce2627efa79f5f # Parent be7565a1115b90b9f8ba8553bbe10fa2634eede6 misc tuning; diff -r be7565a1115b -r 5035a2af185b src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Wed Jun 10 21:49:02 2015 +0200 +++ b/src/HOL/Library/Convex.thy Wed Jun 10 22:28:56 2015 +0200 @@ -3,13 +3,13 @@ Author: Johannes Hoelzl, TU Muenchen *) -section {* Convexity in real vector spaces *} +section \Convexity in real vector spaces\ theory Convex imports Product_Vector begin -subsection {* Convexity. *} +subsection \Convexity\ definition convex :: "'a::real_vector set \ bool" where "convex s \ (\x\s. \y\s. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s)" @@ -57,7 +57,7 @@ lemma convex_UNIV[intro,simp]: "convex UNIV" unfolding convex_def by auto -lemma convex_Inter: "(\s\f. convex s) \ convex(\ f)" +lemma convex_Inter: "(\s\f. convex s) \ convex(\f)" unfolding convex_def by auto lemma convex_Int: "convex s \ convex t \ convex (s \ t)" @@ -103,28 +103,45 @@ and "convex {a..b}" and "convex {a<..b}" and "convex {a.. inner 1 x}" by auto - then show 1: "convex {a..}" by (simp only: convex_halfspace_ge) - have "{..b} = {x. inner 1 x \ b}" by auto - then show 2: "convex {..b}" by (simp only: convex_halfspace_le) - have "{a<..} = {x. a < inner 1 x}" by auto - then show 3: "convex {a<..}" by (simp only: convex_halfspace_gt) - have "{.. {..b}" by auto - then show "convex {a..b}" by (simp only: convex_Int 1 2) - have "{a<..b} = {a<..} \ {..b}" by auto - then show "convex {a<..b}" by (simp only: convex_Int 3 2) - have "{a.. {.. {.. inner 1 x}" + by auto + then show 1: "convex {a..}" + by (simp only: convex_halfspace_ge) + have "{..b} = {x. inner 1 x \ b}" + by auto + then show 2: "convex {..b}" + by (simp only: convex_halfspace_le) + have "{a<..} = {x. a < inner 1 x}" + by auto + then show 3: "convex {a<..}" + by (simp only: convex_halfspace_gt) + have "{.. {..b}" + by auto + then show "convex {a..b}" + by (simp only: convex_Int 1 2) + have "{a<..b} = {a<..} \ {..b}" + by auto + then show "convex {a<..b}" + by (simp only: convex_Int 3 2) + have "{a.. {.. {..Explicit expressions for convexity in terms of arbitrary sums\ lemma convex_setsum: fixes C :: "'a::real_vector set" @@ -151,27 +168,27 @@ have "a i *\<^sub>R y i + (\j\s. a j *\<^sub>R y j) \ C" proof (cases) assume z: "setsum a s = 0" - with `a i + setsum a s = 1` have "a i = 1" + with \a i + setsum a s = 1\ have "a i = 1" by simp - from setsum_nonneg_0 [OF `finite s` _ z] `\j\s. 0 \ a j` have "\j\s. a j = 0" + from setsum_nonneg_0 [OF \finite s\ _ z] \\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 assume nz: "setsum a s \ 0" - with `0 \ setsum a s` have "0 < setsum a s" + with \0 \ setsum a s\ have "0 < setsum a s" by simp then have "(\j\s. (a j / setsum a s) *\<^sub>R y j) \ C" - using `\j\s. 0 \ a j` and `\j\s. y j \ C` + using \\j\s. 0 \ a j\ and \\j\s. y j \ C\ by (simp add: IH setsum_divide_distrib [symmetric]) - from `convex C` and `y i \ C` and this and `0 \ a i` - and `0 \ setsum a s` and `a i + setsum a s = 1` + from \convex C\ and \y i \ C\ and this and \0 \ a i\ + and \0 \ setsum a s\ and \a i + setsum a s = 1\ have "a i *\<^sub>R y i + setsum a s *\<^sub>R (\j\s. (a j / setsum a s) *\<^sub>R y j) \ C" by (rule convexD) then show ?thesis by (simp add: scaleR_setsum_right nz) qed - then show ?case using `finite s` and `i \ s` + then show ?case using \finite s\ and \i \ s\ by simp qed @@ -185,11 +202,10 @@ assume "convex s" "\i. 1 \ i \ i \ k \ 0 \ u i \ x i \ s" "setsum u {1..k} = 1" - from this convex_setsum[of "{1 .. k}" s] - show "(\j\{1 .. k}. u j *\<^sub>R x j) \ s" + with convex_setsum[of "{1 .. k}" s] show "(\j\{1 .. k}. u j *\<^sub>R x j) \ s" by auto next - assume asm: "\k u x. (\ i :: nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 + assume *: "\k u x. (\ i :: nat. 1 \ i \ i \ k \ 0 \ u i \ x i \ s) \ setsum u {1..k} = 1 \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ s" { fix \ :: real @@ -205,7 +221,7 @@ then have "setsum ?u {1 .. 2} = 1" using setsum.If_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] by auto - with asm[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 setsum_head_Suc[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto @@ -213,7 +229,7 @@ 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) + using s by (auto simp: add.commute) } then show "convex s" unfolding convex_alt by auto @@ -233,29 +249,26 @@ then show "(\x\t. u x *\<^sub>R x) \ s" using convex_setsum[of t s u "\ x. x"] by auto next - assume asm0: "\t. \ u. finite t \ t \ s \ (\x\t. 0 \ u x) \ + assume *: "\t. \ u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum 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 asm: "x \ s" "y \ s" "0 \ \" "\ \ 1" - { - assume "x \ y" - then have "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" - using asm0[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] - asm by auto - } - moreover - { - assume "x = y" - then have "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" - using asm0[rule_format, of "{x, y}" "\ z. 1"] - asm by (auto simp: field_simps real_vector.scale_left_diff_distrib) - } - ultimately show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" - by blast + assume **: "x \ s" "y \ s" "0 \ \" "\ \ 1" + show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" + proof (cases "x = y") + case False + then show ?thesis + using *[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] ** + by auto + next + case True + then show ?thesis + using *[rule_format, of "{x, y}" "\ z. 1"] ** + by (auto simp: field_simps real_vector.scale_left_diff_distrib) + qed qed qed @@ -277,7 +290,7 @@ qed (erule_tac x=s in allE, erule_tac x=u in allE, auto) -subsection {* Functions that are convex on a set *} +subsection \Functions that are convex on a set\ definition convex_on :: "'a::real_vector set \ ('a \ real) \ bool" where "convex_on s f \ @@ -299,7 +312,7 @@ assume "0 \ u" "0 \ v" "u + v = 1" ultimately have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" - using assms unfolding convex_on_def by (auto simp add: add_mono) + using assms unfolding convex_on_def by (auto simp: add_mono) then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" by (simp add: field_simps) } @@ -313,7 +326,7 @@ and "convex_on s f" shows "convex_on s (\x. c * f x)" proof - - have *: "\u c fx v fy ::real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" + have *: "\u c fx v fy :: real. u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" by (simp add: field_simps) show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] unfolding convex_on_def and * by auto @@ -330,9 +343,9 @@ proof - let ?m = "max (f x) (f y)" have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" - using assms(4,5) by (auto simp add: mult_left_mono add_mono) + using assms(4,5) by (auto simp: mult_left_mono add_mono) also have "\ = max (f x) (f y)" - using assms(6) by (simp add: distrib_right [symmetric]) + using assms(6) by (simp add: distrib_right [symmetric]) finally show ?thesis using assms unfolding convex_on_def by fastforce qed @@ -340,7 +353,7 @@ lemma convex_on_dist [intro]: fixes s :: "'a::real_normed_vector set" shows "convex_on s (\x. dist a x)" -proof (auto simp add: convex_on_def dist_norm) +proof (auto simp: convex_on_def dist_norm) fix x y assume "x \ s" "y \ s" fix u v :: real @@ -348,16 +361,16 @@ 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 + unfolding scaleR_left_distrib[symmetric] and \u + v = 1\ by simp then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" - by (auto simp add: algebra_simps) + 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 + using \0 \ u\ \0 \ v\ by auto qed -subsection {* Arithmetic operations on sets preserve convexity. *} +subsection \Arithmetic operations on sets preserve convexity\ lemma convex_linear_image: assumes "linear f" @@ -365,7 +378,7 @@ 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 @@ -375,7 +388,7 @@ 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 @@ -386,7 +399,7 @@ 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: @@ -396,7 +409,7 @@ 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) + using \convex s\ by (rule convex_linear_image) qed lemma convex_negations: @@ -406,7 +419,7 @@ have "linear (\x. - x)" by (simp add: linearI) then show ?thesis - using `convex s` by (rule convex_linear_image) + using \convex s\ by (rule convex_linear_image) qed lemma convex_sums: @@ -415,7 +428,7 @@ shows "convex {x + y| x y. x \ s \ y \ t}" proof - have "linear (\(x, y). x + y)" - by (auto intro: linearI simp add: scaleR_add_right) + by (auto intro: linearI simp: scaleR_add_right) with assms have "convex ((\(x, y). x + y) ` (s \ t))" by (intro convex_linear_image convex_Times) also have "((\(x, y). x + y) ` (s \ t)) = {x + y| x y. x \ s \ y \ t}" @@ -428,7 +441,7 @@ shows "convex {x - y| x y. x \ s \ y \ t}" proof - have "{x - y| x y. x \ s \ y \ t} = {x + y |x y. x \ s \ y \ uminus ` t}" - by (auto simp add: diff_conv_add_uminus simp del: add_uminus_conv_diff) + by (auto simp: diff_conv_add_uminus simp del: add_uminus_conv_diff) then show ?thesis using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto qed @@ -457,23 +470,23 @@ unfolding convex_alt proof safe fix y x \ :: real - assume asms: "y > 0" "x > 0" "\ \ 0" "\ \ 1" + assume *: "y > 0" "x > 0" "\ \ 0" "\ \ 1" { assume "\ = 0" then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" by simp - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * by simp } moreover { assume "\ = 1" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms by simp + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * by simp } moreover { assume "\ \ 1" "\ \ 0" - then have "\ > 0" "(1 - \) > 0" using asms by auto - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using asms - by (auto simp add: add_pos_pos) + then have "\ > 0" "(1 - \) > 0" using * by auto + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" using * + by (auto simp: add_pos_pos) } ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" using assms by fastforce @@ -496,70 +509,75 @@ then have ai: "a i = 1" by auto then show ?case by auto next - case (insert i s) note asms = this + case (insert i s) then have "convex_on C f" by simp from this[unfolded convex_on_def, rule_format] have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by simp - { - assume "a i = 1" + show ?case + proof (cases "a i = 1") + case True then have "(\ j \ s. a j) = 0" - using asms by auto + using insert by auto then have "\j. j \ s \ a j = 0" - using setsum_nonneg_0[where 'b=real] asms by fastforce - then have ?case using asms by auto - } - moreover - { - assume asm: "a i \ 1" - from asms have yai: "y i \ C" "a i \ 0" by auto - have fis: "finite (insert i s)" using asms by auto - then have ai1: "a i \ 1" using setsum_nonneg_leq_bound[of "insert i s" a] asms by simp - then have "a i < 1" using asm by auto - then have i0: "1 - a i > 0" by auto + using setsum_nonneg_0[where 'b=real] insert by fastforce + then show ?thesis + using insert by auto + next + case False + from insert have yai: "y i \ C" "a i \ 0" + by auto + have fis: "finite (insert i s)" + using insert by auto + then have ai1: "a i \ 1" + using setsum_nonneg_leq_bound[of "insert i s" a] insert by simp + then have "a i < 1" + using False by auto + then have i0: "1 - a i > 0" + by auto let ?a = "\j. a j / (1 - a i)" - { - fix j - assume "j \ s" - with i0 asms have "?a j \ 0" - by fastforce - } - note a_nonneg = this - have "(\ j \ insert i s. a j) = 1" using asms by auto - then have "(\ j \ s. a j) = 1 - a i" using setsum.insert asms 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 setsum_divide_distrib by simp - have "convex C" using asms by auto + have a_nonneg: "?a j \ 0" if "j \ s" for j + using i0 insert prems by fastforce + have "(\ j \ insert i s. a j) = 1" + using insert by auto + then have "(\ j \ s. a j) = 1 - a i" + using setsum.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 setsum_divide_distrib by simp + have "convex C" using insert by auto then have asum: "(\ j \ s. ?a j *\<^sub>R y j) \ C" - using asms convex_setsum[OF `finite s` - `convex C` a1 a_nonneg] by auto + using insert convex_setsum[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 asms by blast + 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 setsum.insert[of s i "\ j. a j *\<^sub>R y j", OF `finite s` `i \ s`] asms - by (auto simp only:add.commute) + using setsum.insert[of s i "\ j. a j *\<^sub>R y j", OF \finite s\ \i \ s\] insert + by (auto simp only: add.commute) also have "\ = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\ j \ s. a j *\<^sub>R y j) + a i *\<^sub>R y i)" using i0 by auto also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)" using scaleR_right.setsum[of "inverse (1 - a i)" "\ j. a j *\<^sub>R y j" s, symmetric] - by (auto simp:algebra_simps) + by (auto simp: algebra_simps) also have "\ = f ((1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)" by (auto simp: divide_inverse) also have "\ \ (1 - a i) *\<^sub>R f ((\ j \ s. ?a j *\<^sub>R y j)) + a i * f (y i)" using conv[of "y i" "(\ j \ s. ?a j *\<^sub>R y j)" "a i", OF yai(1) asum yai(2) ai1] - by (auto simp add:add.commute) + 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 setsum_right_distrib[of "1 - a i" "\ j. ?a j * f (y j)"] using i0 by auto - also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" using i0 by auto - also have "\ = (\ j \ insert i s. a j * f (y j))" using asms by auto - finally have "f (\ j \ insert i s. a j *\<^sub>R y j) \ (\ j \ insert i s. a j * f (y j))" + also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" + using i0 by auto + also have "\ = (\ j \ insert i s. a j * f (y j))" + using insert by auto + finally show ?thesis by simp - } - ultimately show ?case by auto + qed qed lemma convex_on_alt: @@ -571,24 +589,24 @@ proof safe fix x y fix \ :: real - assume asms: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" + assume *: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" from this[unfolded convex_on_def, rule_format] have "\u v. 0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" by auto - from this[of "\" "1 - \", simplified] asms + from this[of "\" "1 - \", simplified] * show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" by auto next - assume asm: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ + 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 lasm: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" + assume **: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" then have[simp]: "1 - u = v" by auto - from asm[rule_format, of x y u] + from *[rule_format, of x y u] have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" - using lasm by auto + using ** by auto } then show "convex_on C f" unfolding convex_on_def by auto @@ -605,12 +623,12 @@ def a \ "(t - y) / (x - y)" with t have "0 \ a" "0 \ 1 - a" by (auto simp: field_simps) - with f `x \ I` `y \ I` have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" + with f \x \ I\ \y \ I\ have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" by (auto simp: convex_on_def) have "a * x + (1 - a) * y = a * (x - y) + y" by (simp add: field_simps) also have "\ = t" - unfolding a_def using `x < t` `t < y` by simp + unfolding a_def using \x < t\ \t < y\ by simp finally have "f t \ a * f x + (1 - a) * f y" using cvx by simp also have "\ = a * (f x - f y) + f y" @@ -633,17 +651,17 @@ proof safe fix x y \ :: real let ?x = "\ *\<^sub>R x + (1 - \) *\<^sub>R y" - assume asm: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" + assume *: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" then have "1 - \ \ 0" by auto then have xpos: "?x \ C" - using asm unfolding convex_alt by fastforce + using * unfolding convex_alt by fastforce have geq: "\ * (f x - f ?x) + (1 - \) * (f y - f ?x) \ \ * f' ?x * (x - ?x) + (1 - \) * f' ?x * (y - ?x)" - using add_mono[OF mult_left_mono[OF leq[OF xpos asm(2)] `\ \ 0`] - mult_left_mono[OF leq[OF xpos asm(3)] `1 - \ \ 0`]] + using add_mono[OF mult_left_mono[OF leq[OF xpos *(2)] \\ \ 0\] + mult_left_mono[OF leq[OF xpos *(3)] \1 - \ \ 0\]] by auto then have "\ * f x + (1 - \) * f y - f ?x \ 0" - by (auto simp add: field_simps) + by (auto simp: field_simps) then show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" using convex_on_alt by auto qed @@ -654,26 +672,25 @@ and "x \ C" "y \ C" "x < y" shows "{x .. y} \ C" proof safe - fix z assume zasm: "z \ {x .. y}" - { - assume asm: "x < z" "z < y" + fix z assume z: "z \ {x .. y}" + have less: "z \ C" if *: "x < z" "z < y" + proof - let ?\ = "(y - z) / (y - x)" have "0 \ ?\" "?\ \ 1" - using assms asm by (auto simp add: field_simps) + using assms * by (auto simp: field_simps) then have comb: "?\ * x + (1 - ?\) * y \ C" using assms iffD1[OF convex_alt, rule_format, of C y x ?\] by (simp add: algebra_simps) have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" - by (auto simp add: field_simps) + by (auto simp: field_simps) also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" using assms unfolding add_divide_distrib by (auto simp: field_simps) also have "\ = z" using assms by (auto simp: field_simps) - finally have "z \ C" + finally show ?thesis using comb by auto - } - note less = this - show "z \ C" using zasm less assms + qed + show "z \ C" using z less assms unfolding atLeastAtMost_iff le_less by auto qed @@ -689,56 +706,77 @@ proof - { fix x y :: real - assume asm: "x \ C" "y \ C" "y > x" - then have ge: "y - x > 0" "y - x \ 0" by auto - from asm have le: "x - y < 0" "x - y \ 0" by auto + assume *: "x \ C" "y \ C" "y > x" + then have ge: "y - x > 0" "y - x \ 0" + by auto + from * have le: "x - y < 0" "x - y \ 0" + by auto then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" - using subsetD[OF atMostAtLeast_subset_convex[OF `convex C` `x \ C` `y \ C` `x < y`], - THEN f', THEN MVT2[OF `x < y`, rule_format, unfolded atLeastAtMost_iff[symmetric]]] + using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \y \ C\ \x < y\], + THEN f', THEN MVT2[OF \x < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] by auto - then have "z1 \ C" using atMostAtLeast_subset_convex - `convex C` `x \ C` `y \ C` `x < y` by fastforce + 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) + 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 + using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], + THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 by auto obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" - using subsetD[OF atMostAtLeast_subset_convex[OF `convex C` `z1 \ C` `y \ C` `z1 < y`], - THEN f'', THEN MVT2[OF `z1 < y`, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 + 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 asm z1' by auto - also have "\ = (y - z1) * f'' z3" using z3 by auto - finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" by simp - have A': "y - z1 \ 0" using z1 by auto - have "z3 \ C" using z3 asm 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 + using * z1' by auto + also have "\ = (y - z1) * f'' z3" + using z3 by auto + finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" + by simp + have A': "y - z1 \ 0" + using z1 by auto + have "z3 \ C" + using z3 * atMostAtLeast_subset_convex \convex C\ \x \ C\ \z1 \ C\ \x < z1\ + by fastforce + then have B': "f'' z3 \ 0" + using assms by auto + from A' B' have "(y - z1) * f'' z3 \ 0" + by auto + from cool' this have "f' y - (f x - f y) / (x - y) \ 0" + by auto from mult_right_mono_neg[OF this le(2)] have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" by (simp add: algebra_simps) - then have "f' y * (x - y) - (f x - f y) \ 0" using le by auto - then have res: "f' y * (x - y) \ f x - f y" by auto + then have "f' y * (x - y) - (f x - f y) \ 0" + using le by auto + then have res: "f' y * (x - y) \ f x - f y" + by auto have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" - using asm z1 by auto - also have "\ = (z1 - x) * f'' z2" using z2 by auto - finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" by simp - have A: "z1 - x \ 0" using z1 by auto - have "z2 \ C" using z2 z1 asm 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 - from cool this have "(f y - f x) / (y - x) - f' x \ 0" by auto + using * z1 by auto + also have "\ = (z1 - x) * f'' z2" + using z2 by auto + finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" + by simp + have A: "z1 - x \ 0" + using z1 by auto + have "z2 \ C" + using z2 z1 * atMostAtLeast_subset_convex \convex C\ \z1 \ C\ \y \ C\ \z1 < y\ + by fastforce + then have B: "f'' z2 \ 0" + using assms by auto + from A B have "(z1 - x) * f'' z2 \ 0" + by auto + with cool have "(f y - f x) / (y - x) - f' x \ 0" + by auto from mult_right_mono[OF this ge(2)] have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" by (simp add: algebra_simps) - then have "f y - f x - f' x * (y - x) \ 0" using ge by auto + then have "f y - f x - f' x * (y - x) \ 0" + using ge by auto then have "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" - using res by auto } note less_imp = this + using res by auto + } note less_imp = this { fix x y :: real assume "x \ C" "y \ C" "x \ y" @@ -748,7 +786,7 @@ moreover { fix x y :: real - assume asm: "x \ C" "y \ C" "x = y" + assume "x \ C" "y \ C" "x = y" then have "f y - f x \ f' x * (y - x)" by auto } ultimately show ?thesis using assms by blast @@ -781,9 +819,9 @@ by auto then have f''0: "\z::real. z > 0 \ DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" - unfolding inverse_eq_divide by (auto simp add: mult.assoc) + unfolding inverse_eq_divide by (auto simp: mult.assoc) have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" - using `b > 1` by (auto intro!:less_imp_le) + using \b > 1\ by (auto intro!: less_imp_le) from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] show ?thesis by auto