Theory Convex
section ‹Convex Sets and Functions›
theory Convex
imports
Affine
"HOL-Library.Set_Algebras"
begin
subsection ‹Convex Sets›
definition convex :: "'a::real_vector set ⇒ bool"
where "convex s ⟷ (∀x∈s. ∀y∈s. ∀u≥0. ∀v≥0. u + v = 1 ⟶ u *⇩R x + v *⇩R y ∈ s)"
lemma convexI:
assumes "⋀x y u v. x ∈ s ⟹ y ∈ s ⟹ 0 ≤ u ⟹ 0 ≤ v ⟹ u + v = 1 ⟹ u *⇩R x + v *⇩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 *⇩R x + v *⇩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) *⇩R x + u *⇩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 *⇩R x + v *⇩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) *⇩R a + u *⇩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)) *⇩R x + (v/(u+v)) *⇩R y) ∈ S"
using assms
by (simp add: convex_def zero_le_divide_iff add_divide_distrib [symmetric])
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_halfspace_abs_le: "convex {x. ¦inner a x¦ ≤ b}"
proof -
have *: "{x. ¦inner a x¦ ≤ b} = {x. inner a x ≤ b} ∩ {x. -b ≤ inner a x}"
by auto
show ?thesis
unfolding * by (simp add: convex_Int convex_halfspace_ge convex_halfspace_le)
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_halfspace_Re_ge: "convex {x. Re x ≥ b}"
using convex_halfspace_ge[of b "1::complex"] by simp
lemma convex_halfspace_Re_le: "convex {x. Re x ≤ b}"
using convex_halfspace_le[of "1::complex" b] by simp
lemma convex_halfspace_Im_ge: "convex {x. Im x ≥ b}"
using convex_halfspace_ge[of b 𝗂] by simp
lemma convex_halfspace_Im_le: "convex {x. Im x ≤ b}"
using convex_halfspace_le[of 𝗂 b] by simp
lemma convex_halfspace_Re_gt: "convex {x. Re x > b}"
using convex_halfspace_gt[of b "1::complex"] by simp
lemma convex_halfspace_Re_lt: "convex {x. Re x < b}"
using convex_halfspace_lt[of "1::complex" b] by simp
lemma convex_halfspace_Im_gt: "convex {x. Im x > b}"
using convex_halfspace_gt[of b 𝗂] by simp
lemma convex_halfspace_Im_lt: "convex {x. Im x < b}"
using convex_halfspace_lt[of 𝗂 b] by simp
lemma convex_real_interval [iff]:
fixes a b :: "real"
shows "convex {a..}" and "convex {..b}"
and "convex {a<..}" and "convex {..<b}"
and "convex {a..b}" and "convex {a<..b}"
and "convex {a..<b}" and "convex {a<..<b}"
proof -
have "{a..} = {x. 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} = {x. inner 1 x < b}"
by auto
then show 4: "convex {..<b}"
by (simp only: convex_halfspace_lt)
have "{a..b} = {a..} ∩ {..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..<b} = {a..} ∩ {..<b}"
by auto
then show "convex {a..<b}"
by (simp only: convex_Int 1 4)
have "{a<..<b} = {a<..} ∩ {..<b}"
by auto
then show "convex {a<..<b}"
by (simp only: convex_Int 3 4)
qed
lemma convex_Reals: "convex ℝ"
by (simp add: convex_def scaleR_conv_of_real)
subsection ‹Explicit expressions for convexity in terms of arbitrary sums›
lemma convex_sum:
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 *⇩R y j) ∈ C"
using assms(1,3,4,5)
proof (induct arbitrary: a set: finite)
case empty
then show ?case by simp
next
case (insert i S) note IH = this(3)
have "a i + sum a S = 1"
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 ≤ sum a S"
by (simp add: sum_nonneg)
have "a i *⇩R y i + (∑j∈S. a j *⇩R y j) ∈ C"
proof (cases "sum a S = 0")
case True
with ‹a i + sum a S = 1› have "a i = 1"
by simp
from sum_nonneg_0 [OF ‹finite S› _ True] ‹∀j∈S. 0 ≤ a j› have "∀j∈S. a j = 0"
by simp
show ?thesis using ‹a i = 1› and ‹∀j∈S. a j = 0› and ‹y i ∈ C›
by simp
next
case False
with ‹0 ≤ sum a S› have "0 < sum a S"
by simp
then have "(∑j∈S. (a j / sum a S) *⇩R y j) ∈ C"
using ‹∀j∈S. 0 ≤ a j› and ‹∀j∈S. y j ∈ C›
by (simp add: IH sum_divide_distrib [symmetric])
from ‹convex C› and ‹y i ∈ C› and this and ‹0 ≤ a i›
and ‹0 ≤ sum a S› and ‹a i + sum a S = 1›
have "a i *⇩R y i + sum a S *⇩R (∑j∈S. (a j / sum a S) *⇩R y j) ∈ C"
by (rule convexD)
then show ?thesis
by (simp add: scaleR_sum_right False)
qed
then show ?case using ‹finite S› and ‹i ∉ S›
by simp
qed
lemma convex:
"convex S ⟷ (∀(k::nat) u x. (∀i. 1≤i ∧ i≤k ⟶ 0 ≤ u i ∧ x i ∈S) ∧ (sum u {1..k} = 1)
⟶ sum (λi. u i *⇩R x i) {1..k} ∈ S)"
proof safe
fix k :: nat
fix u :: "nat ⇒ real"
fix x
assume "convex S"
"∀i. 1 ≤ i ∧ i ≤ k ⟶ 0 ≤ u i ∧ x i ∈ S"
"sum u {1..k} = 1"
with convex_sum[of "{1 .. k}" S] show "(∑j∈{1 .. k}. u j *⇩R x j) ∈ S"
by auto
next
assume *: "∀k u x. (∀ i :: nat. 1 ≤ i ∧ i ≤ k ⟶ 0 ≤ u i ∧ x i ∈ S) ∧ sum u {1..k} = 1
⟶ (∑i = 1..k. u i *⇩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 "sum ?u {1 .. 2} = 1"
using sum.If_cases[of "{(1 :: nat) .. 2}" "λ x. x = 1" "λ x. μ" "λ x. 1 - μ"]
by auto
with *[rule_format, of "2" ?u ?x] have S: "(∑j ∈ {1..2}. ?u j *⇩R ?x j) ∈ S"
using mu xy by auto
have grarr: "(∑j ∈ {Suc (Suc 0)..2}. ?u j *⇩R ?x j) = (1 - μ) *⇩R y"
using sum.atLeast_Suc_atMost[of "Suc (Suc 0)" 2 "λ j. (1 - μ) *⇩R y"] by auto
from sum.atLeast_Suc_atMost[of "Suc 0" 2 "λ j. ?u j *⇩R ?x j", simplified this]
have "(∑j ∈ {1..2}. ?u j *⇩R ?x j) = μ *⇩R x + (1 - μ) *⇩R y"
by auto
then have "(1 - μ) *⇩R y + μ *⇩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) ∧ sum u t = 1 ⟶ sum (λx. u x *⇩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" "sum u t = 1"
then show "(∑x∈t. u x *⇩R x) ∈ S"
using convex_sum[of t S u "λ x. x"] by auto
next
assume *: "∀t. ∀ u. finite t ∧ t ⊆ S ∧ (∀x∈t. 0 ≤ u x) ∧
sum u t = 1 ⟶ (∑x∈t. u x *⇩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 - μ) *⇩R x + μ *⇩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) ∧ sum u S = 1 ⟶ sum (λx. u x *⇩R x) S ∈ S)"
(is "?lhs = ?rhs")
proof
{ have if_distrib_arg: "⋀P f g x. (if P then f else g) x = (if P then f x else g x)"
by simp
fix T :: "'a set" and u :: "'a ⇒ real"
assume sum: "∀u. (∀x∈S. 0 ≤ u x) ∧ sum u S = 1 ⟶ (∑x∈S. u x *⇩R x) ∈ S"
assume *: "∀x∈T. 0 ≤ u x" "sum u T = 1"
assume "T ⊆ S"
then have "S ∩ T = T" by auto
with sum[THEN spec[where x="λx. if x∈T then u x else 0"]] * have "(∑x∈T. u x *⇩R x) ∈ S"
by (auto simp: assms sum.If_cases if_distrib if_distrib_arg) }
moreover assume ?rhs
ultimately show ?lhs
unfolding convex_explicit by auto
qed (auto simp: convex_explicit assms)
subsection ‹Convex Functions on a Set›
definition 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 *⇩R x + v *⇩R y) ≤ u * f x + v * f y)"
definition concave_on :: "'a::real_vector set ⇒ ('a ⇒ real) ⇒ bool"
where "concave_on S f ≡ convex_on S (λx. - f x)"
lemma concave_on_iff:
"concave_on S f ⟷
(∀x∈S. ∀y∈S. ∀u≥0. ∀v≥0. u + v = 1 ⟶ f (u *⇩R x + v *⇩R y) ≥ u * f x + v * f y)"
by (auto simp: concave_on_def convex_on_def algebra_simps)
lemma convex_onI [intro?]:
assumes "⋀t x y. t > 0 ⟹ t < 1 ⟹ x ∈ A ⟹ y ∈ A ⟹
f ((1 - t) *⇩R x + t *⇩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 *⇩R x + v *⇩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) *⇩R x + t *⇩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) *⇩R x + t *⇩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) *⇩R x + t *⇩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) *⇩R x + t *⇩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 *⇩R x + v *⇩R y) + g (u *⇩R x + v *⇩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 *⇩R x + v *⇩R y) + g (u *⇩R x + v *⇩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 *⇩R x + v *⇩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 *⇩R a + v *⇩R a"
unfolding scaleR_left_distrib[symmetric] and ‹u + v = 1› by simp
then have *: "a - (u *⇩R x + v *⇩R y) = (u *⇩R (a - x)) + (v *⇩R (a - y))"
by (auto simp: algebra_simps)
show "norm (a - (u *⇩R x + v *⇩R y)) ≤ u * norm (a - x) + v * norm (a - y)"
unfolding * using norm_triangle_ineq[of "u *⇩R (a - x)" "v *⇩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 *⇩R x) ` S)"
proof -
have "linear (λx. c *⇩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 *⇩R c) ` S)"
proof -
have "linear (λx. x *⇩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∈ S. ⋃y ∈ T. {x + y})"
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∈ S. ⋃y ∈ T. {x + y})"
by auto
finally show ?thesis .
qed
lemma convex_differences:
assumes "convex S" "convex T"
shows "convex (⋃x∈ S. ⋃y ∈ T. {x - y})"
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:
"convex ((+) a ` S)" if "convex S"
proof -
have "(⋃ x∈ {a}. ⋃y ∈ S. {x + y}) = (+) a ` S"
by auto
then show ?thesis
using convex_sums [OF convex_singleton [of a] that] by auto
qed
lemma convex_translation_subtract:
"convex ((λb. b - a) ` S)" if "convex S"
using convex_translation [of S "- a"] that by (simp cong: image_cong_simp)
lemma convex_affinity:
assumes "convex S"
shows "convex ((λx. a + c *⇩R x) ` S)"
proof -
have "(λx. a + c *⇩R x) ` S = (+) a ` (*⇩R) c ` S"
by auto
then show ?thesis
using convex_translation[OF convex_scaling[OF assms], of a c] by auto
qed
lemma convex_on_sum:
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 *⇩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 (μ *⇩R x + (1 - μ) *⇩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: sum_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 sum_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 sum.insert insert by fastforce
then have "(∑ j ∈ s. a j) / (1 - a i) = 1"
using i0 by auto
then have a1: "(∑ j ∈ s. ?a j) = 1"
unfolding sum_divide_distrib by simp
have "convex C" using insert by auto
then have asum: "(∑ j ∈ s. ?a j *⇩R y j) ∈ C"
using insert convex_sum [OF ‹finite s› ‹convex C› a1 a_nonneg] by auto
have asum_le: "f (∑ j ∈ s. ?a j *⇩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 *⇩R y j) = f ((∑ j ∈ s. a j *⇩R y j) + a i *⇩R y i)"
using sum.insert[of s i "λ j. a j *⇩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)) *⇩R (∑ j ∈ s. a j *⇩R y j) + a i *⇩R y i)"
using i0 by auto
also have "… = f ((1 - a i) *⇩R (∑ j ∈ s. (a j * inverse (1 - a i)) *⇩R y j) + a i *⇩R y i)"
using scaleR_right.sum[of "inverse (1 - a i)" "λ j. a j *⇩R y j" s, symmetric]
by (auto simp: algebra_simps)
also have "… = f ((1 - a i) *⇩R (∑ j ∈ s. ?a j *⇩R y j) + a i *⇩R y i)"
by (auto simp: divide_inverse)
also have "… ≤ (1 - a i) *⇩R f ((∑ j ∈ s. ?a j *⇩R y j)) + a i * f (y i)"
using conv[of "y i" "(∑ j ∈ s. ?a j *⇩R y j)" "a i", OF yai(1) asum yai(2) ai1]
by (auto simp: add.commute)
also have "… ≤ (1 - a i) * (∑ j ∈ s. ?a j * f (y j)) + a i * f (y i)"
using add_right_mono [OF mult_left_mono [of _ _ "1 - a i",
OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"]
by simp
also have "… = (∑ j ∈ s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)"
unfolding sum_distrib_left[of "1 - a i" "λ j. ?a j * f (y j)"]
using i0 by auto
also have "… = (∑ j ∈ s. a j * f (y j)) + a i * f (y i)"
using i0 by auto
also have "… = (∑ j ∈ insert i s. a j * f (y j))"
using insert by auto
finally show ?thesis
by simp
qed
qed
lemma convex_on_alt:
fixes C :: "'a::real_vector set"
shows "convex_on C f ⟷
(∀x ∈ C. ∀ y ∈ C. ∀ μ :: real. μ ≥ 0 ∧ μ ≤ 1 ⟶
f (μ *⇩R x + (1 - μ) *⇩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 *⇩R x + v *⇩R y) ≤ u * f x + v * f y" for u v
by auto
from this [of "μ" "1 - μ", simplified] *
show "f (μ *⇩R x + (1 - μ) *⇩R y) ≤ μ * f x + (1 - μ) * f y"
by auto
next
assume *: "∀x∈C. ∀y∈C. ∀μ. 0 ≤ μ ∧ μ ≤ 1 ⟶
f (μ *⇩R x + (1 - μ) *⇩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 *⇩R x + v *⇩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
using assms
proof safe
fix x y μ :: real
let ?x = "μ *⇩R x + (1 - μ) *⇩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 (μ *⇩R x + (1 - μ) *⇩R y) ≤ μ * f x + (1 - μ) * f y"
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 0: "⋀x. x ∈ C ⟹ f'' x ≥ 0"
shows "convex_on C f"
using f''_imp_f'[OF conv f' f'' 0] assms pos_convex_function
by fastforce
lemma f''_le0_imp_concave:
fixes f :: "real ⇒ real"
assumes "convex C"
and "⋀x. x ∈ C ⟹ DERIV f x :> (f' x)"
and "⋀x. x ∈ C ⟹ DERIV f' x :> (f'' x)"
and "⋀x. x ∈ C ⟹ f'' x ≤ 0"
shows "concave_on C f"
unfolding concave_on_def
by (rule assms f''_ge0_imp_convex derivative_eq_intros | simp)+
lemma log_concave:
fixes b :: real
assumes "b > 1"
shows "concave_on {0<..} (λ x. log b x)"
using assms
by (intro f''_le0_imp_concave derivative_eq_intros | simp)+
lemma ln_concave: "concave_on {0<..} ln"
unfolding log_ln by (simp add: log_concave)
lemma minus_log_convex:
fixes b :: real
assumes "b > 1"
shows "convex_on {0 <..} (λ x. - log b x)"
using assms concave_on_def log_concave by blast
lemma powr_convex:
assumes "p ≥ 1" shows "convex_on {0<..} (λx. x powr p)"
using assms
by (intro f''_ge0_imp_convex derivative_eq_intros | simp)+
lemma exp_convex: "convex_on UNIV exp"
by (intro f''_ge0_imp_convex derivative_eq_intros | simp)+
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) *⇩R x + t *⇩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: field_split_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 field_split_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) *⇩R x + ((c - x) / d) *⇩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 field_split_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)) *⇩R x + (1 - (y - c) / d) *⇩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)
subsection ‹Some inequalities›
lemma Youngs_inequality_0:
fixes a::real
assumes "0 ≤ α" "0 ≤ β" "α+β = 1" "a>0" "b>0"
shows "a powr α * b powr β ≤ α*a + β*b"
proof -
have "α * ln a + β * ln b ≤ ln (α * a + β * b)"
using assms ln_concave by (simp add: concave_on_iff)
moreover have "0 < α * a + β * b"
using assms by (smt (verit) mult_pos_pos split_mult_pos_le)
ultimately show ?thesis
using assms by (simp add: powr_def mult_exp_exp flip: ln_ge_iff)
qed
lemma Youngs_inequality:
fixes p::real
assumes "p>1" "q>1" "1/p + 1/q = 1" "a≥0" "b≥0"
shows "a * b ≤ a powr p / p + b powr q / q"
proof (cases "a=0 ∨ b=0")
case False
then show ?thesis
using Youngs_inequality_0 [of "1/p" "1/q" "a powr p" "b powr q"] assms
by (simp add: powr_powr)
qed (use assms in auto)
lemma Cauchy_Schwarz_ineq_sum:
fixes a :: "'a ⇒ 'b::linordered_field"
shows "(∑i∈I. a i * b i)⇧2 ≤ (∑i∈I. (a i)⇧2) * (∑i∈I. (b i)⇧2)"
proof (cases "(∑i∈I. (b i)⇧2) > 0")
case False
then consider "⋀i. i∈I ⟹ b i = 0" | "infinite I"
by (metis (mono_tags, lifting) sum_pos2 zero_le_power2 zero_less_power2)
thus ?thesis
by fastforce
next
case True
define r where "r ≡ (∑i∈I. a i * b i) / (∑i∈I. (b i)⇧2)"
with True have *: "(∑i∈I. a i * b i) = r * (∑i∈I. (b i)⇧2)"
by simp
have "0 ≤ (∑i∈I. (a i - r * b i)⇧2)"
by (meson sum_nonneg zero_le_power2)
also have "... = (∑i∈I. (a i)⇧2) - 2 * r * (∑i∈I. a i * b i) + r⇧2 * (∑i∈I. (b i)⇧2)"
by (simp add: algebra_simps power2_eq_square sum_distrib_left flip: sum.distrib)
also have "… = (∑i∈I. (a i)⇧2) - (∑i∈I. a i * b i) * r"
by (simp add: * power2_eq_square)
also have "… = (∑i∈I. (a i)⇧2) - ((∑i∈I. a i * b i))⇧2 / (∑i∈I. (b i)⇧2)"
by (simp add: r_def power2_eq_square)
finally have "0 ≤ (∑i∈I. (a i)⇧2) - ((∑i∈I. a i * b i))⇧2 / (∑i∈I. (b i)⇧2)" .
hence "((∑i∈I. a i * b i))⇧2 / (∑i∈I. (b i)⇧2) ≤ (∑i∈I. (a i)⇧2)"
by (simp add: le_diff_eq)
thus "((∑i∈I. a i * b i))⇧2 ≤ (∑i∈I. (a i)⇧2) * (∑i∈I. (b i)⇧2)"
by (simp add: pos_divide_le_eq True)
qed
subsection ‹Misc related lemmas›
lemma convex_translation_eq [simp]:
"convex ((+) a ` s) ⟷ convex s"
by (metis convex_translation translation_galois)
lemma convex_translation_subtract_eq [simp]:
"convex ((λb. b - a) ` s) ⟷ convex s"
using convex_translation_eq [of "- a"] by (simp cong: image_cong_simp)
lemma convex_linear_image_eq [simp]:
fixes f :: "'a::real_vector ⇒ 'b::real_vector"
shows "⟦linear f; inj f⟧ ⟹ convex (f ` s) ⟷ convex s"
by (metis (no_types) convex_linear_image convex_linear_vimage inj_vimage_image_eq)
lemma vector_choose_size:
assumes "0 ≤ c"
obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c"
proof -
obtain a::'a where "a ≠ 0"
using UNIV_not_singleton UNIV_eq_I set_zero singletonI by fastforce
then show ?thesis
by (rule_tac x="scaleR (c / norm a) a" in that) (simp add: assms)
qed
lemma vector_choose_dist:
assumes "0 ≤ c"
obtains y :: "'a::{real_normed_vector, perfect_space}" where "dist x y = c"
by (metis add_diff_cancel_left' assms dist_commute dist_norm vector_choose_size)
lemma sum_delta'':
fixes s::"'a::real_vector set"
assumes "finite s"
shows "(∑x∈s. (if y = x then f x else 0) *⇩R x) = (if y∈s then (f y) *⇩R y else 0)"
proof -
have *: "⋀x y. (if y = x then f x else (0::real)) *⇩R x = (if x=y then (f x) *⇩R x else 0)"
by auto
show ?thesis
unfolding * using sum.delta[OF assms, of y "λx. f x *⇩R x"] by auto
qed
subsection ‹Cones›
definition cone :: "'a::real_vector set ⇒ bool"
where "cone s ⟷ (∀x∈s. ∀c≥0. c *⇩R x ∈ s)"
lemma cone_empty[intro, simp]: "cone {}"
unfolding cone_def by auto
lemma cone_univ[intro, simp]: "cone UNIV"
unfolding cone_def by auto
lemma cone_Inter[intro]: "∀s∈f. cone s ⟹ cone (⋂f)"
unfolding cone_def by auto
lemma subspace_imp_cone: "subspace S ⟹ cone S"
by (simp add: cone_def subspace_scale)
subsubsection ‹Conic hull›
lemma cone_cone_hull: "cone (cone hull S)"
unfolding hull_def by auto
lemma cone_hull_eq: "cone hull S = S ⟷ cone S"
by (metis cone_cone_hull hull_same)
lemma mem_cone:
assumes "cone S" "x ∈ S" "c ≥ 0"
shows "c *⇩R x ∈ S"
using assms cone_def[of S] by auto
lemma cone_contains_0:
assumes "cone S"
shows "S ≠ {} ⟷ 0 ∈ S"
using assms mem_cone by fastforce
lemma cone_0: "cone {0}"
unfolding cone_def by auto
lemma cone_Union[intro]: "(∀s∈f. cone s) ⟶ cone (⋃f)"
unfolding cone_def by blast
lemma cone_iff:
assumes "S ≠ {}"
shows "cone S ⟷ 0 ∈ S ∧ (∀c. c > 0 ⟶ ((*⇩R) c) ` S = S)"
proof -
{
assume "cone S"
{
fix c :: real
assume "c > 0"
{
fix x
assume "x ∈ S"
then have "x ∈ ((*⇩R) c) ` S"
unfolding image_def
using ‹cone S› ‹c>0› mem_cone[of S x "1/c"]
exI[of "(λt. t ∈ S ∧ x = c *⇩R t)" "(1 / c) *⇩R x"]
by auto
}
moreover
{
fix x
assume "x ∈ ((*⇩R) c) ` S"
then have "x ∈ S"
using ‹0 < c› ‹cone S› mem_cone by fastforce
}
ultimately have "((*⇩R) c) ` S = S" by blast
}
then have "0 ∈ S ∧ (∀c. c > 0 ⟶ ((*⇩R) c) ` S = S)"
using ‹cone S› cone_contains_0[of S] assms by auto
}
moreover
{
assume a: "0 ∈ S ∧ (∀c. c > 0 ⟶ ((*⇩R) c) ` S = S)"
{
fix x
assume "x ∈ S"
fix c1 :: real
assume "c1 ≥ 0"
then have "c1 = 0 ∨ c1 > 0" by auto
then have "c1 *⇩R x ∈ S" using a ‹x ∈ S› by auto
}
then have "cone S" unfolding cone_def by auto
}
ultimately show ?thesis by blast
qed
lemma cone_hull_empty: "cone hull {} = {}"
by (metis cone_empty cone_hull_eq)
lemma cone_hull_empty_iff: "S = {} ⟷ cone hull S = {}"
by (metis bot_least cone_hull_empty hull_subset xtrans(5))
lemma cone_hull_contains_0: "S ≠ {} ⟷ 0 ∈ cone hull S"
using cone_cone_hull[of S] cone_contains_0[of "cone hull S"] cone_hull_empty_iff[of S]
by auto
lemma mem_cone_hull:
assumes "x ∈ S" "c ≥ 0"
shows "c *⇩R x ∈ cone hull S"
by (metis assms cone_cone_hull hull_inc mem_cone)
proposition cone_hull_expl: "cone hull S = {c *⇩R x | c x. c ≥ 0 ∧ x ∈ S}"
(is "?lhs = ?rhs")
proof -
{
fix x
assume "x ∈ ?rhs"
then obtain cx :: real and xx where x: "x = cx *⇩R xx" "cx ≥ 0" "xx ∈ S"
by auto
fix c :: real
assume c: "c ≥ 0"
then have "c *⇩R x = (c * cx) *⇩R xx"
using x by (simp add: algebra_simps)
moreover
have "c * cx ≥ 0" using c x by auto
ultimately
have "c *⇩R x ∈ ?rhs" using x by auto
}
then have "cone ?rhs"
unfolding cone_def by auto
then have "?rhs ∈ Collect cone"
unfolding mem_Collect_eq by auto
{
fix x
assume "x ∈ S"
then have "1 *⇩R x ∈ ?rhs"
using zero_le_one by blast
then have "x ∈ ?rhs" by auto
}
then have "S ⊆ ?rhs" by auto
then have "?lhs ⊆ ?rhs"
using ‹?rhs ∈ Collect cone› hull_minimal[of S "?rhs" "cone"] by auto
moreover
{
fix x
assume "x ∈ ?rhs"
then obtain cx :: real and xx where x: "x = cx *⇩R xx" "cx ≥ 0" "xx ∈ S"
by auto
then have "xx ∈ cone hull S"
using hull_subset[of S] by auto
then have "x ∈ ?lhs"
using x cone_cone_hull[of S] cone_def[of "cone hull S"] by auto
}
ultimately show ?thesis by auto
qed
lemma convex_cone:
"convex s ∧ cone s ⟷ (∀x∈s. ∀y∈s. (x + y) ∈ s) ∧ (∀x∈s. ∀c≥0. (c *⇩R x) ∈ s)"
(is "?lhs = ?rhs")
proof -
{
fix x y
assume "x∈s" "y∈s" and ?lhs
then have "2 *⇩R x ∈s" "2 *⇩R y ∈ s"
unfolding cone_def by auto
then have "x + y ∈ s"
using ‹?lhs›[unfolded convex_def, THEN conjunct1]
apply (erule_tac x="2*⇩R x" in ballE)
apply (erule_tac x="2*⇩R y" in ballE)
apply (erule_tac x="1/2" in allE, simp)
apply (erule_tac x="1/2" in allE, auto)
done
}
then show ?thesis
unfolding convex_def cone_def by blast
qed
subsection ‹Connectedness of convex sets›
lemma convex_connected:
fixes S :: "'a::real_normed_vector set"
assumes "convex S"
shows "connected S"
proof (rule connectedI)
fix A B
assume "open A" "open B" "A ∩ B ∩ S = {}" "S ⊆ A ∪ B"
moreover
assume "A ∩ S ≠ {}" "B ∩ S ≠ {}"
then obtain a b where a: "a ∈ A" "a ∈ S" and b: "b ∈ B" "b ∈ S" by auto
define f where [abs_def]: "f u = u *⇩R a + (1 - u) *⇩R b" for u
then have "continuous_on {0 .. 1} f"
by (auto intro!: continuous_intros)
then have "connected (f ` {0 .. 1})"
by (auto intro!: connected_continuous_image)
note connectedD[OF this, of A B]
moreover have "a ∈ A ∩ f ` {0 .. 1}"
using a by (auto intro!: image_eqI[of _ _ 1] simp: f_def)
moreover have "b ∈ B ∩ f ` {0 .. 1}"
using b by (auto intro!: image_eqI[of _ _ 0] simp: f_def)
moreover have "f ` {0 .. 1} ⊆ S"
using ‹convex S› a b unfolding convex_def f_def by auto
ultimately show False by auto
qed
corollary%unimportant connected_UNIV[intro]: "connected (UNIV :: 'a::real_normed_vector set)"
by (simp add: convex_connected)
lemma convex_prod:
assumes "⋀i. i ∈ Basis ⟹ convex {x. P i x}"
shows "convex {x. ∀i∈Basis. P i (x∙i)}"
using assms unfolding convex_def
by (auto simp: inner_add_left)
lemma convex_positive_orthant: "convex {x::'a::euclidean_space. (∀i∈Basis. 0 ≤ x∙i)}"
by (rule convex_prod) (simp flip: atLeast_def)
subsection ‹Convex hull›
lemma convex_convex_hull [iff]: "convex (convex hull s)"
unfolding hull_def
using convex_Inter[of "{t. convex t ∧ s ⊆ t}"]
by auto
lemma convex_hull_subset:
"s ⊆ convex hull t ⟹ convex hull s ⊆ convex hull t"
by (simp add: subset_hull)
lemma convex_hull_eq: "