# HG changeset patch # User hoelzl # Date 1475229617 -7200 # Node ID f4b4fba60b1dff7a94d574d3755f75fe209f4db1 # Parent 4359400adfe7afd29f771ec0bcb2bba795c257f0 HOL-Analysis: move Library/Convex to Convex_Euclidean_Space diff -r 4359400adfe7 -r f4b4fba60b1d src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Fri Sep 30 11:35:39 2016 +0200 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Fri Sep 30 12:00:17 2016 +0200 @@ -137,7 +137,7 @@ moreover have "open ({a<..finite s\ finite_imp_closed) ultimately show "f differentiable at x within {a..b}" - using x le by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) + using x le by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) qed (use x le st dist_real_def in auto) next case ge show ?diff_fg @@ -149,7 +149,7 @@ moreover have "open ({c<..finite t\ finite_imp_closed) ultimately show "g differentiable at x within {a..b}" - using x ge by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) + using x ge by (auto simp add: at_within_open_NO_MATCH differentiable_at_withinI) qed (use x ge st dist_real_def in auto) qed } @@ -1753,7 +1753,7 @@ apply (simp add: segment_convex_hull) apply (rule convex_hull_subset) using assms - apply (auto simp: hull_inc c' Convex.convexD_alt) + apply (auto simp: hull_inc c' convexD_alt) done qed @@ -1770,7 +1770,7 @@ apply (simp_all add: segment_convex_hull) apply (rule_tac [!] convex_hull_subset) using assms - apply (auto simp: hull_inc c' Convex.convexD_alt) + apply (auto simp: hull_inc c' convexD_alt) done show ?thesis apply (rule contour_integral_unique) diff -r 4359400adfe7 -r f4b4fba60b1d src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Fri Sep 30 11:35:39 2016 +0200 +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Fri Sep 30 12:00:17 2016 +0200 @@ -1,5 +1,9 @@ -(* Author: L C Paulson, University of Cambridge - Authors: Robert Himmelmann, TU Muenchen; Bogdan Grechuk, University of Edinburgh +(* Title: HOL/Analysis/Convex_Euclidean_Space.thy + Author: L C Paulson, University of Cambridge + Author: Robert Himmelmann, TU Muenchen + Author: Bogdan Grechuk, University of Edinburgh + Author: Armin Heller, TU Muenchen + Author: Johannes Hoelzl, TU Muenchen *) section \Convex sets, functions and related things\ @@ -7,28 +11,10 @@ theory Convex_Euclidean_Space imports Topology_Euclidean_Space - "~~/src/HOL/Library/Convex" + "~~/src/HOL/Library/Product_Vector" "~~/src/HOL/Library/Set_Algebras" begin -(*FIXME move up e.g. to Library/Convex.thy, but requires the "support" primitive*) -lemma convex_supp_setsum: - assumes "convex S" and 1: "supp_setsum u I = 1" - and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" - shows "supp_setsum (\i. u i *\<^sub>R f i) I \ S" -proof - - have fin: "finite {i \ I. u i \ 0}" - using 1 setsum.infinite by (force simp: supp_setsum_def support_on_def) - then have eq: "supp_setsum (\i. u i *\<^sub>R f i) I = setsum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" - by (force intro: setsum.mono_neutral_left simp: supp_setsum_def support_on_def) - show ?thesis - apply (simp add: eq) - apply (rule convex_setsum [OF fin \convex S\]) - using 1 assms apply (auto simp: supp_setsum_def support_on_def) - done -qed - - lemma dim_image_eq: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes lf: "linear f" @@ -158,9 +144,6 @@ apply auto done -lemma convex_translation_eq [simp]: "convex ((\x. a + x) ` s) \ convex s" - by (metis convex_translation translation_galois) - lemma translation_inverse_subset: assumes "((\x. - a + x) ` V) \ (S :: 'n::ab_group_add set)" shows "V \ ((\x. a + x) ` S)" @@ -179,6 +162,1012 @@ then show ?thesis by auto qed +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)" + +lemma convexI: + assumes "\x y u v. x \ s \ y \ s \ 0 \ u \ 0 \ v \ u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" + shows "convex s" + using assms unfolding convex_def by fast + +lemma convexD: + assumes "convex s" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" + shows "u *\<^sub>R x + v *\<^sub>R y \ s" + using assms unfolding convex_def by fast + +lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R y) \ s)" + (is "_ \ ?alt") +proof + show "convex s" if alt: ?alt + proof - + { + fix x y and u v :: real + assume mem: "x \ s" "y \ s" + assume "0 \ u" "0 \ v" + moreover + assume "u + v = 1" + then have "u = 1 - v" by auto + ultimately have "u *\<^sub>R x + v *\<^sub>R y \ s" + using alt [rule_format, OF mem] by auto + } + then show ?thesis + unfolding convex_def by auto + qed + show ?alt if "convex s" + using that by (auto simp: convex_def) +qed + +lemma convexD_alt: + assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" + shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \ s" + using assms unfolding convex_alt by auto + +lemma mem_convex_alt: + assumes "convex S" "x \ S" "y \ S" "u \ 0" "v \ 0" "u + v > 0" + shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \ S" + apply (rule convexD) + using assms + apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) + done + +lemma convex_empty[intro,simp]: "convex {}" + unfolding convex_def by simp + +lemma convex_singleton[intro,simp]: "convex {a}" + unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) + +lemma convex_UNIV[intro,simp]: "convex UNIV" + unfolding convex_def by auto + +lemma convex_Inter: "(\s. s\f \ convex s) \ convex(\f)" + unfolding convex_def by auto + +lemma convex_Int: "convex s \ convex t \ convex (s \ t)" + unfolding convex_def by auto + +lemma convex_INT: "(\i. i \ A \ convex (B i)) \ convex (\i\A. B i)" + unfolding convex_def by auto + +lemma convex_Times: "convex s \ convex t \ convex (s \ t)" + unfolding convex_def by auto + +lemma convex_halfspace_le: "convex {x. inner a x \ b}" + unfolding convex_def + by (auto simp: inner_add intro!: convex_bound_le) + +lemma convex_halfspace_ge: "convex {x. inner a x \ b}" +proof - + have *: "{x. inner a x \ b} = {x. inner (-a) x \ -b}" + by auto + show ?thesis + unfolding * using convex_halfspace_le[of "-a" "-b"] by auto +qed + +lemma convex_hyperplane: "convex {x. inner a x = b}" +proof - + have *: "{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" + by auto + show ?thesis using convex_halfspace_le convex_halfspace_ge + by (auto intro!: convex_Int simp: *) +qed + +lemma convex_halfspace_lt: "convex {x. inner a x < b}" + unfolding convex_def + by (auto simp: convex_bound_lt inner_add) + +lemma convex_halfspace_gt: "convex {x. inner a x > b}" + using convex_halfspace_lt[of "-a" "-b"] by auto + +lemma convex_real_interval [iff]: + fixes a b :: "real" + shows "convex {a..}" and "convex {..b}" + and "convex {a<..}" and "convex {.. inner 1 x}" + by auto + then show 1: "convex {a..}" + by (simp only: convex_halfspace_ge) + have "{..b} = {x. inner 1 x \ b}" + by auto + then show 2: "convex {..b}" + by (simp only: convex_halfspace_le) + have "{a<..} = {x. a < inner 1 x}" + by auto + then show 3: "convex {a<..}" + by (simp only: convex_halfspace_gt) + have "{.. {..b}" + by auto + then show "convex {a..b}" + by (simp only: convex_Int 1 2) + have "{a<..b} = {a<..} \ {..b}" + by auto + then show "convex {a<..b}" + by (simp only: convex_Int 3 2) + have "{a.. {.. {.." + by (simp add: convex_def scaleR_conv_of_real) + + +subsection \Explicit expressions for convexity in terms of arbitrary sums\ + +lemma convex_setsum: + fixes C :: "'a::real_vector set" + assumes "finite s" + and "convex C" + and "(\ i \ s. a i) = 1" + assumes "\i. i \ s \ a i \ 0" + and "\i. i \ s \ y i \ C" + shows "(\ j \ s. a j *\<^sub>R y j) \ C" + using assms(1,3,4,5) +proof (induct arbitrary: a set: finite) + case empty + then show ?case by simp +next + case (insert i s) note IH = this(3) + have "a i + 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 "setsum a s = 0") + case True + with \a i + setsum a s = 1\ have "a i = 1" + by simp + from setsum_nonneg_0 [OF \finite s\ _ True] \\j\s. 0 \ a j\ have "\j\s. a j = 0" + by simp + show ?thesis using \a i = 1\ and \\j\s. a j = 0\ and \y i \ C\ + by simp + next + case False + with \0 \ 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 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 False) + qed + then show ?case using \finite s\ and \i \ s\ + by simp +qed + +lemma convex: + "convex s \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) + \ setsum (\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" + "setsum u {1..k} = 1" + with convex_setsum[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) \ setsum u {1..k} = 1 + \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ s" + { + fix \ :: real + fix x y :: 'a + assume xy: "x \ s" "y \ s" + assume mu: "\ \ 0" "\ \ 1" + let ?u = "\i. if (i :: nat) = 1 then \ else 1 - \" + let ?x = "\i. if (i :: nat) = 1 then x else y" + have "{1 :: nat .. 2} \ - {x. x = 1} = {2}" + by auto + then have card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" + by simp + then have "setsum ?u {1 .. 2} = 1" + using setsum.If_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] + by auto + with *[rule_format, of "2" ?u ?x] have s: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ s" + using mu xy by auto + have grarr: "(\j \ {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \) *\<^sub>R y" + using setsum_head_Suc[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto + from setsum_head_Suc[of "Suc 0" 2 "\ j. ?u j *\<^sub>R ?x j", simplified this] + have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" + by auto + then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" + using s by (auto simp: add.commute) + } + then show "convex s" + unfolding convex_alt by auto +qed + + +lemma convex_explicit: + fixes s :: "'a::real_vector set" + shows "convex s \ + (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *\<^sub>R x) t \ s)" +proof safe + fix t + fix u :: "'a \ real" + assume "convex s" + and "finite t" + and "t \ s" "\x\t. 0 \ u x" "setsum u t = 1" + then show "(\x\t. u x *\<^sub>R x) \ s" + using convex_setsum[of t s u "\ x. x"] by auto +next + 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 **: "x \ s" "y \ s" "0 \ \" "\ \ 1" + show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" + proof (cases "x = y") + case False + then show ?thesis + using *[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] ** + by auto + next + case True + then show ?thesis + using *[rule_format, of "{x, y}" "\ z. 1"] ** + by (auto simp: field_simps real_vector.scale_left_diff_distrib) + qed + qed +qed + +lemma convex_finite: + assumes "finite s" + shows "convex s \ (\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\x. u x *\<^sub>R x) s \ s)" + unfolding convex_explicit + apply safe + subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto + subgoal for t u + proof - + have if_distrib_arg: "\P f g x. (if P then f else g) x = (if P then f x else g x)" + by simp + assume sum: "\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) \ s" + assume *: "\x\t. 0 \ u x" "setsum 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 setsum.If_cases if_distrib if_distrib_arg) + qed + done + + +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)" + +lemma convex_onI [intro?]: + assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + shows "convex_on A f" + unfolding convex_on_def +proof clarify + fix x y + fix u v :: real + assume A: "x \ A" "y \ A" "u \ 0" "v \ 0" "u + v = 1" + from A(5) have [simp]: "v = 1 - u" + by (simp add: algebra_simps) + from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" + using assms[of u y x] + by (cases "u = 0 \ u = 1") (auto simp: algebra_simps) +qed + +lemma convex_on_linorderI [intro?]: + fixes A :: "('a::{linorder,real_vector}) set" + assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ x < y \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + shows "convex_on A f" +proof + fix x y + fix t :: real + assume A: "x \ A" "y \ A" "t > 0" "t < 1" + with assms [of t x y] assms [of "1 - t" y x] + show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + by (cases x y rule: linorder_cases) (auto simp: algebra_simps) +qed + +lemma convex_onD: + assumes "convex_on A f" + shows "\t x y. t \ 0 \ t \ 1 \ x \ A \ y \ A \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + using assms by (auto simp: convex_on_def) + +lemma convex_onD_Icc: + assumes "convex_on {x..y} f" "x \ (y :: _ :: {real_vector,preorder})" + shows "\t. t \ 0 \ t \ 1 \ + f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" + using assms(2) by (intro convex_onD [OF assms(1)]) simp_all + +lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" + unfolding convex_on_def by auto + +lemma convex_on_add [intro]: + assumes "convex_on s f" + and "convex_on s g" + shows "convex_on s (\x. f x + g x)" +proof - + { + fix x y + assume "x \ s" "y \ s" + moreover + fix u v :: real + assume "0 \ u" "0 \ v" "u + v = 1" + ultimately + have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" + using assms unfolding convex_on_def by (auto simp: add_mono) + then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" + by (simp add: field_simps) + } + then show ?thesis + unfolding convex_on_def by auto +qed + +lemma convex_on_cmul [intro]: + fixes c :: real + assumes "0 \ c" + and "convex_on s f" + shows "convex_on s (\x. c * f x)" +proof - + have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" + for u c fx v fy :: real + by (simp add: field_simps) + show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] + unfolding convex_on_def and * by auto +qed + +lemma convex_lower: + assumes "convex_on s f" + and "x \ s" + and "y \ s" + and "0 \ u" + and "0 \ v" + and "u + v = 1" + shows "f (u *\<^sub>R x + v *\<^sub>R y) \ max (f x) (f y)" +proof - + let ?m = "max (f x) (f y)" + have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" + using assms(4,5) by (auto simp: mult_left_mono add_mono) + also have "\ = max (f x) (f y)" + using assms(6) by (simp add: distrib_right [symmetric]) + finally show ?thesis + using assms unfolding convex_on_def by fastforce +qed + +lemma convex_on_dist [intro]: + fixes s :: "'a::real_normed_vector set" + shows "convex_on s (\x. dist a x)" +proof (auto simp: convex_on_def dist_norm) + fix x y + assume "x \ s" "y \ s" + fix u v :: real + assume "0 \ u" + assume "0 \ v" + assume "u + v = 1" + have "a = u *\<^sub>R a + v *\<^sub>R a" + unfolding scaleR_left_distrib[symmetric] and \u + v = 1\ by simp + then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" + by (auto simp: algebra_simps) + show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" + unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] + using \0 \ u\ \0 \ v\ by auto +qed + + +subsection \Arithmetic operations on sets preserve convexity\ + +lemma convex_linear_image: + assumes "linear f" + and "convex s" + shows "convex (f ` s)" +proof - + interpret f: linear f by fact + from \convex s\ show "convex (f ` s)" + by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) +qed + +lemma convex_linear_vimage: + assumes "linear f" + and "convex s" + shows "convex (f -` s)" +proof - + interpret f: linear f by fact + from \convex s\ show "convex (f -` s)" + by (simp add: convex_def f.add f.scaleR) +qed + +lemma convex_scaling: + assumes "convex s" + shows "convex ((\x. c *\<^sub>R x) ` s)" +proof - + have "linear (\x. c *\<^sub>R x)" + by (simp add: linearI scaleR_add_right) + then show ?thesis + using \convex s\ by (rule convex_linear_image) +qed + +lemma convex_scaled: + assumes "convex s" + shows "convex ((\x. x *\<^sub>R c) ` s)" +proof - + have "linear (\x. x *\<^sub>R c)" + by (simp add: linearI scaleR_add_left) + then show ?thesis + using \convex s\ by (rule convex_linear_image) +qed + +lemma convex_negations: + assumes "convex s" + shows "convex ((\x. - x) ` s)" +proof - + have "linear (\x. - x)" + by (simp add: linearI) + then show ?thesis + using \convex s\ by (rule convex_linear_image) +qed + +lemma convex_sums: + assumes "convex s" + and "convex t" + shows "convex {x + y| x y. x \ s \ y \ t}" +proof - + have "linear (\(x, y). x + y)" + by (auto intro: linearI simp: scaleR_add_right) + with assms have "convex ((\(x, y). x + y) ` (s \ t))" + by (intro convex_linear_image convex_Times) + also have "((\(x, y). x + y) ` (s \ t)) = {x + y| x y. x \ s \ y \ t}" + by auto + finally show ?thesis . +qed + +lemma convex_differences: + assumes "convex s" "convex t" + 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: diff_conv_add_uminus simp del: add_uminus_conv_diff) + then show ?thesis + using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto +qed + +lemma convex_translation: + assumes "convex s" + shows "convex ((\x. a + x) ` s)" +proof - + have "{a + y |y. y \ s} = (\x. a + x) ` s" + by auto + then show ?thesis + using convex_sums[OF convex_singleton[of a] assms] by auto +qed + +lemma convex_affinity: + assumes "convex s" + shows "convex ((\x. a + c *\<^sub>R x) ` s)" +proof - + have "(\x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" + by auto + then show ?thesis + using convex_translation[OF convex_scaling[OF assms], of a c] by auto +qed + +lemma pos_is_convex: "convex {0 :: real <..}" + unfolding convex_alt +proof safe + fix y x \ :: real + assume *: "y > 0" "x > 0" "\ \ 0" "\ \ 1" + { + assume "\ = 0" + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" + by simp + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" + using * by simp + } + moreover + { + assume "\ = 1" + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" + using * by simp + } + moreover + { + assume "\ \ 1" "\ \ 0" + then have "\ > 0" "(1 - \) > 0" + using * by auto + then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" + using * by (auto simp: add_pos_pos) + } + ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" + by fastforce +qed + +lemma convex_on_setsum: + fixes a :: "'a \ real" + and y :: "'a \ 'b::real_vector" + and f :: "'b \ real" + assumes "finite s" "s \ {}" + and "convex_on C f" + and "convex C" + and "(\ i \ s. a i) = 1" + and "\i. i \ s \ a i \ 0" + and "\i. i \ s \ y i \ C" + shows "f (\ i \ s. a i *\<^sub>R y i) \ (\ i \ s. a i * f (y i))" + using assms +proof (induct s arbitrary: a rule: finite_ne_induct) + case (singleton i) + then have ai: "a i = 1" + by auto + then show ?case + by auto +next + case (insert i s) + then have "convex_on C f" + by simp + from this[unfolded convex_on_def, rule_format] + have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ + f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + by simp + show ?case + proof (cases "a i = 1") + case True + then have "(\ j \ s. a j) = 0" + using insert by auto + then have "\j. j \ s \ a j = 0" + using insert by (fastforce simp: setsum_nonneg_eq_0_iff) + then show ?thesis + using insert by auto + next + case False + from insert have yai: "y i \ C" "a i \ 0" + by auto + have fis: "finite (insert i s)" + using insert by auto + then have ai1: "a i \ 1" + using 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)" + have a_nonneg: "?a j \ 0" if "j \ s" for j + using i0 insert that by fastforce + have "(\ j \ insert i s. a j) = 1" + using insert by auto + then have "(\ j \ s. a j) = 1 - a i" + using 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 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 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\] 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) + 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 setsum_distrib_left[of "1 - a i" "\ j. ?a j * f (y j)"] + using i0 by auto + also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" + using i0 by auto + also have "\ = (\ j \ insert i s. a j * f (y j))" + using insert by auto + finally show ?thesis + by simp + qed +qed + +lemma convex_on_alt: + fixes C :: "'a::real_vector set" + assumes "convex C" + shows "convex_on C f \ + (\x \ C. \ y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ + f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" +proof safe + fix x y + fix \ :: real + assume *: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" + from this[unfolded convex_on_def, rule_format] + have "0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" for u v + by auto + from this [of "\" "1 - \", simplified] * + show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + by auto +next + assume *: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ + f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + { + fix x y + fix u v :: real + assume **: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" + then have[simp]: "1 - u = v" by auto + from *[rule_format, of x y u] + have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" + using ** by auto + } + then show "convex_on C f" + unfolding convex_on_def by auto +qed + +lemma convex_on_diff: + fixes f :: "real \ real" + assumes f: "convex_on I f" + and I: "x \ I" "y \ I" + and t: "x < t" "t < y" + shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + and "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" +proof - + define a where "a \ (t - y) / (x - y)" + with t have "0 \ a" "0 \ 1 - a" + by (auto simp: field_simps) + with f \x \ I\ \y \ I\ have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" + by (auto simp: convex_on_def) + have "a * x + (1 - a) * y = a * (x - y) + y" + by (simp add: field_simps) + also have "\ = t" + unfolding a_def using \x < t\ \t < y\ by simp + finally have "f t \ a * f x + (1 - a) * f y" + using cvx by simp + also have "\ = a * (f x - f y) + f y" + by (simp add: field_simps) + finally have "f t - f y \ a * (f x - f y)" + by simp + with t show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + by (simp add: le_divide_eq divide_le_eq field_simps a_def) + with t show "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" + by (simp add: le_divide_eq divide_le_eq field_simps) +qed + +lemma pos_convex_function: + fixes f :: "real \ real" + assumes "convex C" + and leq: "\x y. x \ C \ y \ C \ f' x * (y - x) \ f y - f x" + shows "convex_on C f" + unfolding convex_on_alt[OF assms(1)] + using assms +proof safe + fix x y \ :: real + let ?x = "\ *\<^sub>R x + (1 - \) *\<^sub>R y" + assume *: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" + then have "1 - \ \ 0" by auto + then have xpos: "?x \ C" + using * unfolding convex_alt by fastforce + have geq: "\ * (f x - f ?x) + (1 - \) * (f y - f ?x) \ + \ * f' ?x * (x - ?x) + (1 - \) * f' ?x * (y - ?x)" + using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \\ \ 0\] + mult_left_mono [OF leq [OF xpos *(3)] \1 - \ \ 0\]] + by auto + then have "\ * f x + (1 - \) * f y - f ?x \ 0" + by (auto simp: field_simps) + then show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" + using convex_on_alt by auto +qed + +lemma atMostAtLeast_subset_convex: + fixes C :: "real set" + assumes "convex C" + and "x \ C" "y \ C" "x < y" + shows "{x .. y} \ C" +proof safe + fix z assume z: "z \ {x .. y}" + have less: "z \ C" if *: "x < z" "z < y" + proof - + let ?\ = "(y - z) / (y - x)" + have "0 \ ?\" "?\ \ 1" + using assms * by (auto simp: field_simps) + then have comb: "?\ * x + (1 - ?\) * y \ C" + using assms iffD1[OF convex_alt, rule_format, of C y x ?\] + by (simp add: algebra_simps) + have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" + by (auto simp: field_simps) + also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" + using assms by (simp only: add_divide_distrib) (auto simp: field_simps) + also have "\ = z" + using assms by (auto simp: field_simps) + finally show ?thesis + using comb by auto + qed + show "z \ C" + using z less assms by (auto simp: le_less) +qed + +lemma f''_imp_f': + fixes f :: "real \ real" + assumes "convex C" + and f': "\x. x \ C \ DERIV f x :> (f' x)" + and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" + and pos: "\x. x \ C \ f'' x \ 0" + and x: "x \ C" + and y: "y \ C" + shows "f' x * (y - x) \ f y - f x" + using assms +proof - + have less_imp: "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" + if *: "x \ C" "y \ C" "y > x" for x y :: real + proof - + from * have ge: "y - x > 0" "y - x \ 0" + by auto + from * have le: "x - y < 0" "x - y \ 0" + by auto + then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" + using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \y \ C\ \x < y\], + THEN f', THEN MVT2[OF \x < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] + by auto + then have "z1 \ C" + using atMostAtLeast_subset_convex \convex C\ \x \ C\ \y \ C\ \x < y\ + by fastforce + from z1 have z1': "f x - f y = (x - y) * f' z1" + by (simp add: field_simps) + obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" + using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], + THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 + by auto + obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" + using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \z1 \ C\ \y \ C\ \z1 < y\], + THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 + by auto + have "f' y - (f x - f y) / (x - y) = f' y - f' z1" + using * z1' by auto + also have "\ = (y - z1) * f'' z3" + using z3 by auto + finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" + by simp + have A': "y - z1 \ 0" + using z1 by auto + have "z3 \ C" + using z3 * atMostAtLeast_subset_convex \convex C\ \x \ C\ \z1 \ C\ \x < z1\ + by fastforce + then have B': "f'' z3 \ 0" + using assms by auto + from A' B' have "(y - z1) * f'' z3 \ 0" + by auto + from cool' this have "f' y - (f x - f y) / (x - y) \ 0" + by auto + from mult_right_mono_neg[OF this le(2)] + have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" + by (simp add: algebra_simps) + then have "f' y * (x - y) - (f x - f y) \ 0" + using le by auto + then have res: "f' y * (x - y) \ f x - f y" + by auto + have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" + using * z1 by auto + also have "\ = (z1 - x) * f'' z2" + using z2 by auto + finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" + by simp + have A: "z1 - x \ 0" + using z1 by auto + have "z2 \ C" + using z2 z1 * atMostAtLeast_subset_convex \convex C\ \z1 \ C\ \y \ C\ \z1 < y\ + by fastforce + then have B: "f'' z2 \ 0" + using assms by auto + from A B have "(z1 - x) * f'' z2 \ 0" + by auto + with cool have "(f y - f x) / (y - x) - f' x \ 0" + by auto + from mult_right_mono[OF this ge(2)] + have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" + by (simp add: algebra_simps) + then have "f y - f x - f' x * (y - x) \ 0" + using ge by auto + then show "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" + using res by auto + qed + show ?thesis + proof (cases "x = y") + case True + with x y show ?thesis by auto + next + case False + with less_imp x y show ?thesis + by (auto simp: neq_iff) + qed +qed + +lemma f''_ge0_imp_convex: + fixes f :: "real \ real" + assumes conv: "convex C" + and f': "\x. x \ C \ DERIV f x :> (f' x)" + and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" + and pos: "\x. x \ C \ f'' x \ 0" + shows "convex_on C f" + using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function + by fastforce + +lemma minus_log_convex: + fixes b :: real + assumes "b > 1" + shows "convex_on {0 <..} (\ x. - log b x)" +proof - + have "\z. z > 0 \ DERIV (log b) z :> 1 / (ln b * z)" + using DERIV_log by auto + then have f': "\z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" + by (auto simp: DERIV_minus) + have "\z::real. z > 0 \ DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" + using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto + from this[THEN DERIV_cmult, of _ "- 1 / ln b"] + have "\z::real. z > 0 \ + DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" + by auto + then have f''0: "\z::real. z > 0 \ + DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" + unfolding inverse_eq_divide by (auto simp: mult.assoc) + have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" + using \b > 1\ by (auto intro!: less_imp_le) + from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] + show ?thesis + by auto +qed + + +subsection \Convexity of real functions\ + +lemma convex_on_realI: + assumes "connected A" + and "\x. x \ A \ (f has_real_derivative f' x) (at x)" + and "\x y. x \ A \ y \ A \ x \ y \ f' x \ f' y" + shows "convex_on A f" +proof (rule convex_on_linorderI) + fix t x y :: real + assume t: "t > 0" "t < 1" + assume xy: "x \ A" "y \ A" "x < y" + define z where "z = (1 - t) * x + t * y" + with \connected A\ and xy have ivl: "{x..y} \ A" + using connected_contains_Icc by blast + + from xy t have xz: "z > x" + by (simp add: z_def algebra_simps) + have "y - z = (1 - t) * (y - x)" + by (simp add: z_def algebra_simps) + also from xy t have "\ > 0" + by (intro mult_pos_pos) simp_all + finally have yz: "z < y" + by simp + + from assms xz yz ivl t have "\\. \ > x \ \ < z \ f z - f x = (z - x) * f' \" + by (intro MVT2) (auto intro!: assms(2)) + then obtain \ where \: "\ > x" "\ < z" "f' \ = (f z - f x) / (z - x)" + by auto + from assms xz yz ivl t have "\\. \ > z \ \ < y \ f y - f z = (y - z) * f' \" + by (intro MVT2) (auto intro!: assms(2)) + then obtain \ where \: "\ > z" "\ < y" "f' \ = (f y - f z) / (y - z)" + by auto + + from \(3) have "(f y - f z) / (y - z) = f' \" .. + also from \ \ ivl have "\ \ A" "\ \ A" + by auto + with \ \ have "f' \ \ f' \" + by (intro assms(3)) auto + also from \(3) have "f' \ = (f z - f x) / (z - x)" . + finally have "(f y - f z) * (z - x) \ (f z - f x) * (y - z)" + using xz yz by (simp add: field_simps) + also have "z - x = t * (y - x)" + by (simp add: z_def algebra_simps) + also have "y - z = (1 - t) * (y - x)" + by (simp add: z_def algebra_simps) + finally have "(f y - f z) * t \ (f z - f x) * (1 - t)" + using xy by simp + then show "(1 - t) * f x + t * f y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" + by (simp add: z_def algebra_simps) +qed + +lemma convex_on_inverse: + assumes "A \ {0<..}" + shows "convex_on A (inverse :: real \ real)" +proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\x. -inverse (x^2)"]) + fix u v :: real + assume "u \ {0<..}" "v \ {0<..}" "u \ v" + with assms show "-inverse (u^2) \ -inverse (v^2)" + by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) +qed (insert assms, auto intro!: derivative_eq_intros simp: divide_simps power2_eq_square) + +lemma convex_onD_Icc': + assumes "convex_on {x..y} f" "c \ {x..y}" + defines "d \ y - x" + shows "f c \ (f y - f x) / d * (c - x) + f x" +proof (cases x y rule: linorder_cases) + case less + then have d: "d > 0" + by (simp add: d_def) + from assms(2) less have A: "0 \ (c - x) / d" "(c - x) / d \ 1" + by (simp_all add: d_def divide_simps) + have "f c = f (x + (c - x) * 1)" + by simp + also from less have "1 = ((y - x) / d)" + by (simp add: d_def) + also from d have "x + (c - x) * \ = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" + by (simp add: field_simps) + also have "f \ \ (1 - (c - x) / d) * f x + (c - x) / d * f y" + using assms less by (intro convex_onD_Icc) simp_all + also from d have "\ = (f y - f x) / d * (c - x) + f x" + by (simp add: field_simps) + finally show ?thesis . +qed (insert assms(2), simp_all) + +lemma convex_onD_Icc'': + assumes "convex_on {x..y} f" "c \ {x..y}" + defines "d \ y - x" + shows "f c \ (f x - f y) / d * (y - c) + f y" +proof (cases x y rule: linorder_cases) + case less + then have d: "d > 0" + by (simp add: d_def) + from assms(2) less have A: "0 \ (y - c) / d" "(y - c) / d \ 1" + by (simp_all add: d_def divide_simps) + have "f c = f (y - (y - c) * 1)" + by simp + also from less have "1 = ((y - x) / d)" + by (simp add: d_def) + also from d have "y - (y - c) * \ = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" + by (simp add: field_simps) + also have "f \ \ (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" + using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) + also from d have "\ = (f x - f y) / d * (y - c) + f y" + by (simp add: field_simps) + finally show ?thesis . +qed (insert assms(2), simp_all) + +lemma convex_supp_setsum: + assumes "convex S" and 1: "supp_setsum u I = 1" + and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" + shows "supp_setsum (\i. u i *\<^sub>R f i) I \ S" +proof - + have fin: "finite {i \ I. u i \ 0}" + using 1 setsum.infinite by (force simp: supp_setsum_def support_on_def) + then have eq: "supp_setsum (\i. u i *\<^sub>R f i) I = setsum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" + by (force intro: setsum.mono_neutral_left simp: supp_setsum_def support_on_def) + show ?thesis + apply (simp add: eq) + apply (rule convex_setsum [OF fin \convex S\]) + using 1 assms apply (auto simp: supp_setsum_def support_on_def) + done +qed + +lemma convex_translation_eq [simp]: "convex ((\x. a + x) ` s) \ convex s" + by (metis convex_translation translation_galois) + lemma convex_linear_image_eq [simp]: fixes f :: "'a::real_vector \ 'b::real_vector" shows "\linear f; inj f\ \ convex (f ` s) \ convex s" diff -r 4359400adfe7 -r f4b4fba60b1d src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Fri Sep 30 11:35:39 2016 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,999 +0,0 @@ -(* Title: HOL/Library/Convex.thy - Author: Armin Heller, TU Muenchen - Author: Johannes Hoelzl, TU Muenchen -*) - -section \Convexity in real vector spaces\ - -theory Convex - imports Product_Vector -begin - -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)" - -lemma convexI: - assumes "\x y u v. x \ s \ y \ s \ 0 \ u \ 0 \ v \ u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \ s" - shows "convex s" - using assms unfolding convex_def by fast - -lemma convexD: - assumes "convex s" and "x \ s" and "y \ s" and "0 \ u" and "0 \ v" and "u + v = 1" - shows "u *\<^sub>R x + v *\<^sub>R y \ s" - using assms unfolding convex_def by fast - -lemma convex_alt: "convex s \ (\x\s. \y\s. \u. 0 \ u \ u \ 1 \ ((1 - u) *\<^sub>R x + u *\<^sub>R y) \ s)" - (is "_ \ ?alt") -proof - show "convex s" if alt: ?alt - proof - - { - fix x y and u v :: real - assume mem: "x \ s" "y \ s" - assume "0 \ u" "0 \ v" - moreover - assume "u + v = 1" - then have "u = 1 - v" by auto - ultimately have "u *\<^sub>R x + v *\<^sub>R y \ s" - using alt [rule_format, OF mem] by auto - } - then show ?thesis - unfolding convex_def by auto - qed - show ?alt if "convex s" - using that by (auto simp: convex_def) -qed - -lemma convexD_alt: - assumes "convex s" "a \ s" "b \ s" "0 \ u" "u \ 1" - shows "((1 - u) *\<^sub>R a + u *\<^sub>R b) \ s" - using assms unfolding convex_alt by auto - -lemma mem_convex_alt: - assumes "convex S" "x \ S" "y \ S" "u \ 0" "v \ 0" "u + v > 0" - shows "((u/(u+v)) *\<^sub>R x + (v/(u+v)) *\<^sub>R y) \ S" - apply (rule convexD) - using assms - apply (simp_all add: zero_le_divide_iff add_divide_distrib [symmetric]) - done - -lemma convex_empty[intro,simp]: "convex {}" - unfolding convex_def by simp - -lemma convex_singleton[intro,simp]: "convex {a}" - unfolding convex_def by (auto simp: scaleR_left_distrib[symmetric]) - -lemma convex_UNIV[intro,simp]: "convex UNIV" - unfolding convex_def by auto - -lemma convex_Inter: "(\s. s\f \ convex s) \ convex(\f)" - unfolding convex_def by auto - -lemma convex_Int: "convex s \ convex t \ convex (s \ t)" - unfolding convex_def by auto - -lemma convex_INT: "(\i. i \ A \ convex (B i)) \ convex (\i\A. B i)" - unfolding convex_def by auto - -lemma convex_Times: "convex s \ convex t \ convex (s \ t)" - unfolding convex_def by auto - -lemma convex_halfspace_le: "convex {x. inner a x \ b}" - unfolding convex_def - by (auto simp: inner_add intro!: convex_bound_le) - -lemma convex_halfspace_ge: "convex {x. inner a x \ b}" -proof - - have *: "{x. inner a x \ b} = {x. inner (-a) x \ -b}" - by auto - show ?thesis - unfolding * using convex_halfspace_le[of "-a" "-b"] by auto -qed - -lemma convex_hyperplane: "convex {x. inner a x = b}" -proof - - have *: "{x. inner a x = b} = {x. inner a x \ b} \ {x. inner a x \ b}" - by auto - show ?thesis using convex_halfspace_le convex_halfspace_ge - by (auto intro!: convex_Int simp: *) -qed - -lemma convex_halfspace_lt: "convex {x. inner a x < b}" - unfolding convex_def - by (auto simp: convex_bound_lt inner_add) - -lemma convex_halfspace_gt: "convex {x. inner a x > b}" - using convex_halfspace_lt[of "-a" "-b"] by auto - -lemma convex_real_interval [iff]: - fixes a b :: "real" - shows "convex {a..}" and "convex {..b}" - and "convex {a<..}" and "convex {.. inner 1 x}" - by auto - then show 1: "convex {a..}" - by (simp only: convex_halfspace_ge) - have "{..b} = {x. inner 1 x \ b}" - by auto - then show 2: "convex {..b}" - by (simp only: convex_halfspace_le) - have "{a<..} = {x. a < inner 1 x}" - by auto - then show 3: "convex {a<..}" - by (simp only: convex_halfspace_gt) - have "{.. {..b}" - by auto - then show "convex {a..b}" - by (simp only: convex_Int 1 2) - have "{a<..b} = {a<..} \ {..b}" - by auto - then show "convex {a<..b}" - by (simp only: convex_Int 3 2) - have "{a.. {.. {.." - by (simp add: convex_def scaleR_conv_of_real) - - -subsection \Explicit expressions for convexity in terms of arbitrary sums\ - -lemma convex_setsum: - fixes C :: "'a::real_vector set" - assumes "finite s" - and "convex C" - and "(\ i \ s. a i) = 1" - assumes "\i. i \ s \ a i \ 0" - and "\i. i \ s \ y i \ C" - shows "(\ j \ s. a j *\<^sub>R y j) \ C" - using assms(1,3,4,5) -proof (induct arbitrary: a set: finite) - case empty - then show ?case by simp -next - case (insert i s) note IH = this(3) - have "a i + 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 "setsum a s = 0") - case True - with \a i + setsum a s = 1\ have "a i = 1" - by simp - from setsum_nonneg_0 [OF \finite s\ _ True] \\j\s. 0 \ a j\ have "\j\s. a j = 0" - by simp - show ?thesis using \a i = 1\ and \\j\s. a j = 0\ and \y i \ C\ - by simp - next - case False - with \0 \ 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 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 False) - qed - then show ?case using \finite s\ and \i \ s\ - by simp -qed - -lemma convex: - "convex s \ (\(k::nat) u x. (\i. 1\i \ i\k \ 0 \ u i \ x i \s) \ (setsum u {1..k} = 1) - \ setsum (\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" - "setsum u {1..k} = 1" - with convex_setsum[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) \ setsum u {1..k} = 1 - \ (\i = 1..k. u i *\<^sub>R (x i :: 'a)) \ s" - { - fix \ :: real - fix x y :: 'a - assume xy: "x \ s" "y \ s" - assume mu: "\ \ 0" "\ \ 1" - let ?u = "\i. if (i :: nat) = 1 then \ else 1 - \" - let ?x = "\i. if (i :: nat) = 1 then x else y" - have "{1 :: nat .. 2} \ - {x. x = 1} = {2}" - by auto - then have card: "card ({1 :: nat .. 2} \ - {x. x = 1}) = 1" - by simp - then have "setsum ?u {1 .. 2} = 1" - using setsum.If_cases[of "{(1 :: nat) .. 2}" "\ x. x = 1" "\ x. \" "\ x. 1 - \"] - by auto - with *[rule_format, of "2" ?u ?x] have s: "(\j \ {1..2}. ?u j *\<^sub>R ?x j) \ s" - using mu xy by auto - have grarr: "(\j \ {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \) *\<^sub>R y" - using setsum_head_Suc[of "Suc (Suc 0)" 2 "\ j. (1 - \) *\<^sub>R y"] by auto - from setsum_head_Suc[of "Suc 0" 2 "\ j. ?u j *\<^sub>R ?x j", simplified this] - have "(\j \ {1..2}. ?u j *\<^sub>R ?x j) = \ *\<^sub>R x + (1 - \) *\<^sub>R y" - by auto - then have "(1 - \) *\<^sub>R y + \ *\<^sub>R x \ s" - using s by (auto simp: add.commute) - } - then show "convex s" - unfolding convex_alt by auto -qed - - -lemma convex_explicit: - fixes s :: "'a::real_vector set" - shows "convex s \ - (\t u. finite t \ t \ s \ (\x\t. 0 \ u x) \ setsum u t = 1 \ setsum (\x. u x *\<^sub>R x) t \ s)" -proof safe - fix t - fix u :: "'a \ real" - assume "convex s" - and "finite t" - and "t \ s" "\x\t. 0 \ u x" "setsum u t = 1" - then show "(\x\t. u x *\<^sub>R x) \ s" - using convex_setsum[of t s u "\ x. x"] by auto -next - 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 **: "x \ s" "y \ s" "0 \ \" "\ \ 1" - show "(1 - \) *\<^sub>R x + \ *\<^sub>R y \ s" - proof (cases "x = y") - case False - then show ?thesis - using *[rule_format, of "{x, y}" "\ z. if z = x then 1 - \ else \"] ** - by auto - next - case True - then show ?thesis - using *[rule_format, of "{x, y}" "\ z. 1"] ** - by (auto simp: field_simps real_vector.scale_left_diff_distrib) - qed - qed -qed - -lemma convex_finite: - assumes "finite s" - shows "convex s \ (\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ setsum (\x. u x *\<^sub>R x) s \ s)" - unfolding convex_explicit - apply safe - subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto - subgoal for t u - proof - - have if_distrib_arg: "\P f g x. (if P then f else g) x = (if P then f x else g x)" - by simp - assume sum: "\u. (\x\s. 0 \ u x) \ setsum u s = 1 \ (\x\s. u x *\<^sub>R x) \ s" - assume *: "\x\t. 0 \ u x" "setsum 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 setsum.If_cases if_distrib if_distrib_arg) - qed - done - - -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)" - -lemma convex_onI [intro?]: - assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - shows "convex_on A f" - unfolding convex_on_def -proof clarify - fix x y - fix u v :: real - assume A: "x \ A" "y \ A" "u \ 0" "v \ 0" "u + v = 1" - from A(5) have [simp]: "v = 1 - u" - by (simp add: algebra_simps) - from A(1-4) show "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" - using assms[of u y x] - by (cases "u = 0 \ u = 1") (auto simp: algebra_simps) -qed - -lemma convex_on_linorderI [intro?]: - fixes A :: "('a::{linorder,real_vector}) set" - assumes "\t x y. t > 0 \ t < 1 \ x \ A \ y \ A \ x < y \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - shows "convex_on A f" -proof - fix x y - fix t :: real - assume A: "x \ A" "y \ A" "t > 0" "t < 1" - with assms [of t x y] assms [of "1 - t" y x] - show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - by (cases x y rule: linorder_cases) (auto simp: algebra_simps) -qed - -lemma convex_onD: - assumes "convex_on A f" - shows "\t x y. t \ 0 \ t \ 1 \ x \ A \ y \ A \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - using assms by (auto simp: convex_on_def) - -lemma convex_onD_Icc: - assumes "convex_on {x..y} f" "x \ (y :: _ :: {real_vector,preorder})" - shows "\t. t \ 0 \ t \ 1 \ - f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" - using assms(2) by (intro convex_onD [OF assms(1)]) simp_all - -lemma convex_on_subset: "convex_on t f \ s \ t \ convex_on s f" - unfolding convex_on_def by auto - -lemma convex_on_add [intro]: - assumes "convex_on s f" - and "convex_on s g" - shows "convex_on s (\x. f x + g x)" -proof - - { - fix x y - assume "x \ s" "y \ s" - moreover - fix u v :: real - assume "0 \ u" "0 \ v" "u + v = 1" - ultimately - have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ (u * f x + v * f y) + (u * g x + v * g y)" - using assms unfolding convex_on_def by (auto simp: add_mono) - then have "f (u *\<^sub>R x + v *\<^sub>R y) + g (u *\<^sub>R x + v *\<^sub>R y) \ u * (f x + g x) + v * (f y + g y)" - by (simp add: field_simps) - } - then show ?thesis - unfolding convex_on_def by auto -qed - -lemma convex_on_cmul [intro]: - fixes c :: real - assumes "0 \ c" - and "convex_on s f" - shows "convex_on s (\x. c * f x)" -proof - - have *: "u * (c * fx) + v * (c * fy) = c * (u * fx + v * fy)" - for u c fx v fy :: real - by (simp add: field_simps) - show ?thesis using assms(2) and mult_left_mono [OF _ assms(1)] - unfolding convex_on_def and * by auto -qed - -lemma convex_lower: - assumes "convex_on s f" - and "x \ s" - and "y \ s" - and "0 \ u" - and "0 \ v" - and "u + v = 1" - shows "f (u *\<^sub>R x + v *\<^sub>R y) \ max (f x) (f y)" -proof - - let ?m = "max (f x) (f y)" - have "u * f x + v * f y \ u * max (f x) (f y) + v * max (f x) (f y)" - using assms(4,5) by (auto simp: mult_left_mono add_mono) - also have "\ = max (f x) (f y)" - using assms(6) by (simp add: distrib_right [symmetric]) - finally show ?thesis - using assms unfolding convex_on_def by fastforce -qed - -lemma convex_on_dist [intro]: - fixes s :: "'a::real_normed_vector set" - shows "convex_on s (\x. dist a x)" -proof (auto simp: convex_on_def dist_norm) - fix x y - assume "x \ s" "y \ s" - fix u v :: real - assume "0 \ u" - assume "0 \ v" - assume "u + v = 1" - have "a = u *\<^sub>R a + v *\<^sub>R a" - unfolding scaleR_left_distrib[symmetric] and \u + v = 1\ by simp - then have *: "a - (u *\<^sub>R x + v *\<^sub>R y) = (u *\<^sub>R (a - x)) + (v *\<^sub>R (a - y))" - by (auto simp: algebra_simps) - show "norm (a - (u *\<^sub>R x + v *\<^sub>R y)) \ u * norm (a - x) + v * norm (a - y)" - unfolding * using norm_triangle_ineq[of "u *\<^sub>R (a - x)" "v *\<^sub>R (a - y)"] - using \0 \ u\ \0 \ v\ by auto -qed - - -subsection \Arithmetic operations on sets preserve convexity\ - -lemma convex_linear_image: - assumes "linear f" - and "convex s" - shows "convex (f ` s)" -proof - - interpret f: linear f by fact - from \convex s\ show "convex (f ` s)" - by (simp add: convex_def f.scaleR [symmetric] f.add [symmetric]) -qed - -lemma convex_linear_vimage: - assumes "linear f" - and "convex s" - shows "convex (f -` s)" -proof - - interpret f: linear f by fact - from \convex s\ show "convex (f -` s)" - by (simp add: convex_def f.add f.scaleR) -qed - -lemma convex_scaling: - assumes "convex s" - shows "convex ((\x. c *\<^sub>R x) ` s)" -proof - - have "linear (\x. c *\<^sub>R x)" - by (simp add: linearI scaleR_add_right) - then show ?thesis - using \convex s\ by (rule convex_linear_image) -qed - -lemma convex_scaled: - assumes "convex s" - shows "convex ((\x. x *\<^sub>R c) ` s)" -proof - - have "linear (\x. x *\<^sub>R c)" - by (simp add: linearI scaleR_add_left) - then show ?thesis - using \convex s\ by (rule convex_linear_image) -qed - -lemma convex_negations: - assumes "convex s" - shows "convex ((\x. - x) ` s)" -proof - - have "linear (\x. - x)" - by (simp add: linearI) - then show ?thesis - using \convex s\ by (rule convex_linear_image) -qed - -lemma convex_sums: - assumes "convex s" - and "convex t" - shows "convex {x + y| x y. x \ s \ y \ t}" -proof - - have "linear (\(x, y). x + y)" - by (auto intro: linearI simp: scaleR_add_right) - with assms have "convex ((\(x, y). x + y) ` (s \ t))" - by (intro convex_linear_image convex_Times) - also have "((\(x, y). x + y) ` (s \ t)) = {x + y| x y. x \ s \ y \ t}" - by auto - finally show ?thesis . -qed - -lemma convex_differences: - assumes "convex s" "convex t" - 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: diff_conv_add_uminus simp del: add_uminus_conv_diff) - then show ?thesis - using convex_sums[OF assms(1) convex_negations[OF assms(2)]] by auto -qed - -lemma convex_translation: - assumes "convex s" - shows "convex ((\x. a + x) ` s)" -proof - - have "{a + y |y. y \ s} = (\x. a + x) ` s" - by auto - then show ?thesis - using convex_sums[OF convex_singleton[of a] assms] by auto -qed - -lemma convex_affinity: - assumes "convex s" - shows "convex ((\x. a + c *\<^sub>R x) ` s)" -proof - - have "(\x. a + c *\<^sub>R x) ` s = op + a ` op *\<^sub>R c ` s" - by auto - then show ?thesis - using convex_translation[OF convex_scaling[OF assms], of a c] by auto -qed - -lemma pos_is_convex: "convex {0 :: real <..}" - unfolding convex_alt -proof safe - fix y x \ :: real - assume *: "y > 0" "x > 0" "\ \ 0" "\ \ 1" - { - assume "\ = 0" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y = y" - by simp - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by simp - } - moreover - { - assume "\ = 1" - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by simp - } - moreover - { - assume "\ \ 1" "\ \ 0" - then have "\ > 0" "(1 - \) > 0" - using * by auto - then have "\ *\<^sub>R x + (1 - \) *\<^sub>R y > 0" - using * by (auto simp: add_pos_pos) - } - ultimately show "(1 - \) *\<^sub>R y + \ *\<^sub>R x > 0" - by fastforce -qed - -lemma convex_on_setsum: - fixes a :: "'a \ real" - and y :: "'a \ 'b::real_vector" - and f :: "'b \ real" - assumes "finite s" "s \ {}" - and "convex_on C f" - and "convex C" - and "(\ i \ s. a i) = 1" - and "\i. i \ s \ a i \ 0" - and "\i. i \ s \ y i \ C" - shows "f (\ i \ s. a i *\<^sub>R y i) \ (\ i \ s. a i * f (y i))" - using assms -proof (induct s arbitrary: a rule: finite_ne_induct) - case (singleton i) - then have ai: "a i = 1" - by auto - then show ?case - by auto -next - case (insert i s) - then have "convex_on C f" - by simp - from this[unfolded convex_on_def, rule_format] - have conv: "\x y \. x \ C \ y \ C \ 0 \ \ \ \ \ 1 \ - f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - by simp - show ?case - proof (cases "a i = 1") - case True - then have "(\ j \ s. a j) = 0" - using insert by auto - then have "\j. j \ s \ a j = 0" - using insert by (fastforce simp: setsum_nonneg_eq_0_iff) - then show ?thesis - using insert by auto - next - case False - from insert have yai: "y i \ C" "a i \ 0" - by auto - have fis: "finite (insert i s)" - using insert by auto - then have ai1: "a i \ 1" - using 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)" - have a_nonneg: "?a j \ 0" if "j \ s" for j - using i0 insert that by fastforce - have "(\ j \ insert i s. a j) = 1" - using insert by auto - then have "(\ j \ s. a j) = 1 - a i" - using 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 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 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\] 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) - 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 setsum_distrib_left[of "1 - a i" "\ j. ?a j * f (y j)"] - using i0 by auto - also have "\ = (\ j \ s. a j * f (y j)) + a i * f (y i)" - using i0 by auto - also have "\ = (\ j \ insert i s. a j * f (y j))" - using insert by auto - finally show ?thesis - by simp - qed -qed - -lemma convex_on_alt: - fixes C :: "'a::real_vector set" - assumes "convex C" - shows "convex_on C f \ - (\x \ C. \ y \ C. \ \ :: real. \ \ 0 \ \ \ 1 \ - f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y)" -proof safe - fix x y - fix \ :: real - assume *: "convex_on C f" "x \ C" "y \ C" "0 \ \" "\ \ 1" - from this[unfolded convex_on_def, rule_format] - have "0 \ u \ 0 \ v \ u + v = 1 \ f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" for u v - by auto - from this [of "\" "1 - \", simplified] * - show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - by auto -next - assume *: "\x\C. \y\C. \\. 0 \ \ \ \ \ 1 \ - f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - { - fix x y - fix u v :: real - assume **: "x \ C" "y \ C" "u \ 0" "v \ 0" "u + v = 1" - then have[simp]: "1 - u = v" by auto - from *[rule_format, of x y u] - have "f (u *\<^sub>R x + v *\<^sub>R y) \ u * f x + v * f y" - using ** by auto - } - then show "convex_on C f" - unfolding convex_on_def by auto -qed - -lemma convex_on_diff: - fixes f :: "real \ real" - assumes f: "convex_on I f" - and I: "x \ I" "y \ I" - and t: "x < t" "t < y" - shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" - and "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" -proof - - define a where "a \ (t - y) / (x - y)" - with t have "0 \ a" "0 \ 1 - a" - by (auto simp: field_simps) - with f \x \ I\ \y \ I\ have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" - by (auto simp: convex_on_def) - have "a * x + (1 - a) * y = a * (x - y) + y" - by (simp add: field_simps) - also have "\ = t" - unfolding a_def using \x < t\ \t < y\ by simp - finally have "f t \ a * f x + (1 - a) * f y" - using cvx by simp - also have "\ = a * (f x - f y) + f y" - by (simp add: field_simps) - finally have "f t - f y \ a * (f x - f y)" - by simp - with t show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" - by (simp add: le_divide_eq divide_le_eq field_simps a_def) - with t show "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" - by (simp add: le_divide_eq divide_le_eq field_simps) -qed - -lemma pos_convex_function: - fixes f :: "real \ real" - assumes "convex C" - and leq: "\x y. x \ C \ y \ C \ f' x * (y - x) \ f y - f x" - shows "convex_on C f" - unfolding convex_on_alt[OF assms(1)] - using assms -proof safe - fix x y \ :: real - let ?x = "\ *\<^sub>R x + (1 - \) *\<^sub>R y" - assume *: "convex C" "x \ C" "y \ C" "\ \ 0" "\ \ 1" - then have "1 - \ \ 0" by auto - then have xpos: "?x \ C" - using * unfolding convex_alt by fastforce - have geq: "\ * (f x - f ?x) + (1 - \) * (f y - f ?x) \ - \ * f' ?x * (x - ?x) + (1 - \) * f' ?x * (y - ?x)" - using add_mono [OF mult_left_mono [OF leq [OF xpos *(2)] \\ \ 0\] - mult_left_mono [OF leq [OF xpos *(3)] \1 - \ \ 0\]] - by auto - then have "\ * f x + (1 - \) * f y - f ?x \ 0" - by (auto simp: field_simps) - then show "f (\ *\<^sub>R x + (1 - \) *\<^sub>R y) \ \ * f x + (1 - \) * f y" - using convex_on_alt by auto -qed - -lemma atMostAtLeast_subset_convex: - fixes C :: "real set" - assumes "convex C" - and "x \ C" "y \ C" "x < y" - shows "{x .. y} \ C" -proof safe - fix z assume z: "z \ {x .. y}" - have less: "z \ C" if *: "x < z" "z < y" - proof - - let ?\ = "(y - z) / (y - x)" - have "0 \ ?\" "?\ \ 1" - using assms * by (auto simp: field_simps) - then have comb: "?\ * x + (1 - ?\) * y \ C" - using assms iffD1[OF convex_alt, rule_format, of C y x ?\] - by (simp add: algebra_simps) - have "?\ * x + (1 - ?\) * y = (y - z) * x / (y - x) + (1 - (y - z) / (y - x)) * y" - by (auto simp: field_simps) - also have "\ = ((y - z) * x + (y - x - (y - z)) * y) / (y - x)" - using assms by (simp only: add_divide_distrib) (auto simp: field_simps) - also have "\ = z" - using assms by (auto simp: field_simps) - finally show ?thesis - using comb by auto - qed - show "z \ C" - using z less assms by (auto simp: le_less) -qed - -lemma f''_imp_f': - fixes f :: "real \ real" - assumes "convex C" - and f': "\x. x \ C \ DERIV f x :> (f' x)" - and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" - and pos: "\x. x \ C \ f'' x \ 0" - and x: "x \ C" - and y: "y \ C" - shows "f' x * (y - x) \ f y - f x" - using assms -proof - - have less_imp: "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" - if *: "x \ C" "y \ C" "y > x" for x y :: real - proof - - from * have ge: "y - x > 0" "y - x \ 0" - by auto - from * have le: "x - y < 0" "x - y \ 0" - by auto - then obtain z1 where z1: "z1 > x" "z1 < y" "f y - f x = (y - x) * f' z1" - using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \y \ C\ \x < y\], - THEN f', THEN MVT2[OF \x < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] - by auto - then have "z1 \ C" - using atMostAtLeast_subset_convex \convex C\ \x \ C\ \y \ C\ \x < y\ - by fastforce - from z1 have z1': "f x - f y = (x - y) * f' z1" - by (simp add: field_simps) - obtain z2 where z2: "z2 > x" "z2 < z1" "f' z1 - f' x = (z1 - x) * f'' z2" - using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \x \ C\ \z1 \ C\ \x < z1\], - THEN f'', THEN MVT2[OF \x < z1\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 - by auto - obtain z3 where z3: "z3 > z1" "z3 < y" "f' y - f' z1 = (y - z1) * f'' z3" - using subsetD[OF atMostAtLeast_subset_convex[OF \convex C\ \z1 \ C\ \y \ C\ \z1 < y\], - THEN f'', THEN MVT2[OF \z1 < y\, rule_format, unfolded atLeastAtMost_iff[symmetric]]] z1 - by auto - have "f' y - (f x - f y) / (x - y) = f' y - f' z1" - using * z1' by auto - also have "\ = (y - z1) * f'' z3" - using z3 by auto - finally have cool': "f' y - (f x - f y) / (x - y) = (y - z1) * f'' z3" - by simp - have A': "y - z1 \ 0" - using z1 by auto - have "z3 \ C" - using z3 * atMostAtLeast_subset_convex \convex C\ \x \ C\ \z1 \ C\ \x < z1\ - by fastforce - then have B': "f'' z3 \ 0" - using assms by auto - from A' B' have "(y - z1) * f'' z3 \ 0" - by auto - from cool' this have "f' y - (f x - f y) / (x - y) \ 0" - by auto - from mult_right_mono_neg[OF this le(2)] - have "f' y * (x - y) - (f x - f y) / (x - y) * (x - y) \ 0 * (x - y)" - by (simp add: algebra_simps) - then have "f' y * (x - y) - (f x - f y) \ 0" - using le by auto - then have res: "f' y * (x - y) \ f x - f y" - by auto - have "(f y - f x) / (y - x) - f' x = f' z1 - f' x" - using * z1 by auto - also have "\ = (z1 - x) * f'' z2" - using z2 by auto - finally have cool: "(f y - f x) / (y - x) - f' x = (z1 - x) * f'' z2" - by simp - have A: "z1 - x \ 0" - using z1 by auto - have "z2 \ C" - using z2 z1 * atMostAtLeast_subset_convex \convex C\ \z1 \ C\ \y \ C\ \z1 < y\ - by fastforce - then have B: "f'' z2 \ 0" - using assms by auto - from A B have "(z1 - x) * f'' z2 \ 0" - by auto - with cool have "(f y - f x) / (y - x) - f' x \ 0" - by auto - from mult_right_mono[OF this ge(2)] - have "(f y - f x) / (y - x) * (y - x) - f' x * (y - x) \ 0 * (y - x)" - by (simp add: algebra_simps) - then have "f y - f x - f' x * (y - x) \ 0" - using ge by auto - then show "f y - f x \ f' x * (y - x)" "f' y * (x - y) \ f x - f y" - using res by auto - qed - show ?thesis - proof (cases "x = y") - case True - with x y show ?thesis by auto - next - case False - with less_imp x y show ?thesis - by (auto simp: neq_iff) - qed -qed - -lemma f''_ge0_imp_convex: - fixes f :: "real \ real" - assumes conv: "convex C" - and f': "\x. x \ C \ DERIV f x :> (f' x)" - and f'': "\x. x \ C \ DERIV f' x :> (f'' x)" - and pos: "\x. x \ C \ f'' x \ 0" - shows "convex_on C f" - using f''_imp_f'[OF conv f' f'' pos] assms pos_convex_function - by fastforce - -lemma minus_log_convex: - fixes b :: real - assumes "b > 1" - shows "convex_on {0 <..} (\ x. - log b x)" -proof - - have "\z. z > 0 \ DERIV (log b) z :> 1 / (ln b * z)" - using DERIV_log by auto - then have f': "\z. z > 0 \ DERIV (\ z. - log b z) z :> - 1 / (ln b * z)" - by (auto simp: DERIV_minus) - have "\z::real. z > 0 \ DERIV inverse z :> - (inverse z ^ Suc (Suc 0))" - using less_imp_neq[THEN not_sym, THEN DERIV_inverse] by auto - from this[THEN DERIV_cmult, of _ "- 1 / ln b"] - have "\z::real. z > 0 \ - DERIV (\ z. (- 1 / ln b) * inverse z) z :> (- 1 / ln b) * (- (inverse z ^ Suc (Suc 0)))" - by auto - then have f''0: "\z::real. z > 0 \ - DERIV (\ z. - 1 / (ln b * z)) z :> 1 / (ln b * z * z)" - unfolding inverse_eq_divide by (auto simp: mult.assoc) - have f''_ge0: "\z::real. z > 0 \ 1 / (ln b * z * z) \ 0" - using \b > 1\ by (auto intro!: less_imp_le) - from f''_ge0_imp_convex[OF pos_is_convex, unfolded greaterThan_iff, OF f' f''0 f''_ge0] - show ?thesis - by auto -qed - - -subsection \Convexity of real functions\ - -lemma convex_on_realI: - assumes "connected A" - and "\x. x \ A \ (f has_real_derivative f' x) (at x)" - and "\x y. x \ A \ y \ A \ x \ y \ f' x \ f' y" - shows "convex_on A f" -proof (rule convex_on_linorderI) - fix t x y :: real - assume t: "t > 0" "t < 1" - assume xy: "x \ A" "y \ A" "x < y" - define z where "z = (1 - t) * x + t * y" - with \connected A\ and xy have ivl: "{x..y} \ A" - using connected_contains_Icc by blast - - from xy t have xz: "z > x" - by (simp add: z_def algebra_simps) - have "y - z = (1 - t) * (y - x)" - by (simp add: z_def algebra_simps) - also from xy t have "\ > 0" - by (intro mult_pos_pos) simp_all - finally have yz: "z < y" - by simp - - from assms xz yz ivl t have "\\. \ > x \ \ < z \ f z - f x = (z - x) * f' \" - by (intro MVT2) (auto intro!: assms(2)) - then obtain \ where \: "\ > x" "\ < z" "f' \ = (f z - f x) / (z - x)" - by auto - from assms xz yz ivl t have "\\. \ > z \ \ < y \ f y - f z = (y - z) * f' \" - by (intro MVT2) (auto intro!: assms(2)) - then obtain \ where \: "\ > z" "\ < y" "f' \ = (f y - f z) / (y - z)" - by auto - - from \(3) have "(f y - f z) / (y - z) = f' \" .. - also from \ \ ivl have "\ \ A" "\ \ A" - by auto - with \ \ have "f' \ \ f' \" - by (intro assms(3)) auto - also from \(3) have "f' \ = (f z - f x) / (z - x)" . - finally have "(f y - f z) * (z - x) \ (f z - f x) * (y - z)" - using xz yz by (simp add: field_simps) - also have "z - x = t * (y - x)" - by (simp add: z_def algebra_simps) - also have "y - z = (1 - t) * (y - x)" - by (simp add: z_def algebra_simps) - finally have "(f y - f z) * t \ (f z - f x) * (1 - t)" - using xy by simp - then show "(1 - t) * f x + t * f y \ f ((1 - t) *\<^sub>R x + t *\<^sub>R y)" - by (simp add: z_def algebra_simps) -qed - -lemma convex_on_inverse: - assumes "A \ {0<..}" - shows "convex_on A (inverse :: real \ real)" -proof (rule convex_on_subset[OF _ assms], intro convex_on_realI[of _ _ "\x. -inverse (x^2)"]) - fix u v :: real - assume "u \ {0<..}" "v \ {0<..}" "u \ v" - with assms show "-inverse (u^2) \ -inverse (v^2)" - by (intro le_imp_neg_le le_imp_inverse_le power_mono) (simp_all) -qed (insert assms, auto intro!: derivative_eq_intros simp: divide_simps power2_eq_square) - -lemma convex_onD_Icc': - assumes "convex_on {x..y} f" "c \ {x..y}" - defines "d \ y - x" - shows "f c \ (f y - f x) / d * (c - x) + f x" -proof (cases x y rule: linorder_cases) - case less - then have d: "d > 0" - by (simp add: d_def) - from assms(2) less have A: "0 \ (c - x) / d" "(c - x) / d \ 1" - by (simp_all add: d_def divide_simps) - have "f c = f (x + (c - x) * 1)" - by simp - also from less have "1 = ((y - x) / d)" - by (simp add: d_def) - also from d have "x + (c - x) * \ = (1 - (c - x) / d) *\<^sub>R x + ((c - x) / d) *\<^sub>R y" - by (simp add: field_simps) - also have "f \ \ (1 - (c - x) / d) * f x + (c - x) / d * f y" - using assms less by (intro convex_onD_Icc) simp_all - also from d have "\ = (f y - f x) / d * (c - x) + f x" - by (simp add: field_simps) - finally show ?thesis . -qed (insert assms(2), simp_all) - -lemma convex_onD_Icc'': - assumes "convex_on {x..y} f" "c \ {x..y}" - defines "d \ y - x" - shows "f c \ (f x - f y) / d * (y - c) + f y" -proof (cases x y rule: linorder_cases) - case less - then have d: "d > 0" - by (simp add: d_def) - from assms(2) less have A: "0 \ (y - c) / d" "(y - c) / d \ 1" - by (simp_all add: d_def divide_simps) - have "f c = f (y - (y - c) * 1)" - by simp - also from less have "1 = ((y - x) / d)" - by (simp add: d_def) - also from d have "y - (y - c) * \ = (1 - (1 - (y - c) / d)) *\<^sub>R x + (1 - (y - c) / d) *\<^sub>R y" - by (simp add: field_simps) - also have "f \ \ (1 - (1 - (y - c) / d)) * f x + (1 - (y - c) / d) * f y" - using assms less by (intro convex_onD_Icc) (simp_all add: field_simps) - also from d have "\ = (f x - f y) / d * (y - c) + f y" - by (simp add: field_simps) - finally show ?thesis . -qed (insert assms(2), simp_all) - -end diff -r 4359400adfe7 -r f4b4fba60b1d src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Fri Sep 30 11:35:39 2016 +0200 +++ b/src/HOL/Library/Library.thy Fri Sep 30 12:00:17 2016 +0200 @@ -11,7 +11,6 @@ Char_ord Code_Test Continuum_Not_Denumerable - Convex Combine_PER Complete_Partial_Order2 Countable