diff -r e6d570cb0f5e -r df6133adb63f src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Tue Mar 04 22:30:12 2014 +0100 +++ b/src/HOL/Library/Convex.thy Tue Mar 04 14:00:59 2014 -0800 @@ -111,7 +111,6 @@ then show "convex {a<.. 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 -proof (induct s arbitrary:a rule: finite_induct) + using assms(1,3,4,5) +proof (induct arbitrary: a set: finite) case empty - then show ?case by auto + then show ?case by simp next - case (insert i s) note asms = this - { assume "a i = 1" - then have "(\ j \ s. a j) = 0" - using asms 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 1] asms by simp - then have "a i < 1" using asm by auto - then have i0: "1 - a i > 0" by auto - let ?a = "\j. a j / (1 - a i)" - { fix j assume "j \ s" - then have "?a j \ 0" - using i0 asms divide_nonneg_pos - 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 - with asms have "(\j\s. ?a j *\<^sub>R y j) \ C" using a_nonneg by fastforce - then have "a i *\<^sub>R y i + (1 - a i) *\<^sub>R (\ j \ s. ?a j *\<^sub>R y j) \ C" - using asms yai ai1 by (auto intro: convexD) - then have "a i *\<^sub>R y i + (\ j \ s. (1 - a i) *\<^sub>R (?a j *\<^sub>R y j)) \ C" - using scaleR_right.setsum[of "(1 - a i)" "\ j. ?a j *\<^sub>R y j" s] by auto - then have "a i *\<^sub>R y i + (\ j \ s. a j *\<^sub>R y j) \ C" using i0 by auto - then have ?case using setsum.insert asms by auto - } - ultimately show ?case by auto + case (insert i s) note IH = this(3) + have "a i + setsum a s = 1" and "0 \ a i" and "\j\s. 0 \ a j" and "y i \ C" and "\j\s. y j \ C" + using insert.hyps(1,2) insert.prems by simp_all + then have "0 \ setsum a s" by (simp add: setsum_nonneg) + 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" by simp + 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` by simp + next + assume nz: "setsum a s \ 0" + 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` + by (simp add: IH divide_nonneg_pos 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` + 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` by simp qed lemma convex: @@ -247,6 +236,8 @@ by (auto simp: assms setsum_cases if_distrib if_distrib_arg) qed (erule_tac x=s in allE, erule_tac x=u in allE, auto) +subsection {* Functions that are convex on a set *} + definition 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)"