Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit f1f4c23

Browse files
committed
feat(analysis/convex/basic): convex_on lemmas (#7933)
1 parent 5d03dcd commit f1f4c23

File tree

1 file changed

+81
-22
lines changed

1 file changed

+81
-22
lines changed

src/analysis/convex/basic.lean

Lines changed: 81 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -844,31 +844,29 @@ lemma concave_on.smul [ordered_module ℝ β] {f : E → β} {c : ℝ} (hc : 0
844844
(hf : concave_on s f) : concave_on s (λx, c • f x) :=
845845
@convex_on.smul _ _ _ _ (order_dual β) _ _ _ f c hc hf
846846

847+
section linear_order
848+
variables {γ : Type*} [linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ]
849+
{f : E → γ}
850+
847851
/-- A convex function on a segment is upper-bounded by the max of its endpoints. -/
848-
lemma convex_on.le_on_segment' {γ : Type*}
849-
[linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ]
850-
{f : E → γ} {x y : E} {a b : ℝ}
851-
(hf : convex_on s f) (hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) :
852+
lemma convex_on.le_on_segment' (hf : convex_on s f) {x y : E} {a b : ℝ}
853+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) :
852854
f (a • x + b • y) ≤ max (f x) (f y) :=
853855
calc
854856
f (a • x + b • y) ≤ a • f x + b • f y : hf.2 hx hy ha hb hab
855857
... ≤ a • max (f x) (f y) + b • max (f x) (f y) :
856858
add_le_add (smul_le_smul_of_nonneg (le_max_left _ _) ha)
857859
(smul_le_smul_of_nonneg (le_max_right _ _) hb)
858-
... max (f x) (f y) : by rw [←add_smul, hab, one_smul]
860+
... = max (f x) (f y) : by rw [←add_smul, hab, one_smul]
859861

860862
/-- A concave function on a segment is lower-bounded by the min of its endpoints. -/
861-
lemma concave_on.le_on_segment' {γ : Type*}
862-
[linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ]
863-
{f : E → γ} {x y : E} {a b : ℝ}
864-
(hf : concave_on s f) (hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) :
863+
lemma concave_on.le_on_segment' (hf : concave_on s f) {x y : E} {a b : ℝ}
864+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 ≤ b) (hab : a + b = 1) :
865865
min (f x) (f y) ≤ f (a • x + b • y) :=
866-
@convex_on.le_on_segment' _ _ _ _ (order_dual γ) _ _ _ f x y a b hf hx hy ha hb hab
866+
@convex_on.le_on_segment' _ _ _ _ (order_dual γ) _ _ _ f hf x y a b hx hy ha hb hab
867867

868868
/-- A convex function on a segment is upper-bounded by the max of its endpoints. -/
869-
lemma convex_on.le_on_segment {γ : Type*}
870-
[linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ]
871-
{f : E → γ} (hf : convex_on s f) {x y z : E}
869+
lemma convex_on.le_on_segment (hf : convex_on s f) {x y z : E}
872870
(hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ [x, y]) :
873871
f z ≤ max (f x) (f y) :=
874872
let ⟨a, b, ha, hb, hab, hz⟩ := hz in hz ▸ hf.le_on_segment' hx hy ha hb hab
@@ -881,18 +879,80 @@ lemma concave_on.le_on_segment {γ : Type*}
881879
min (f x) (f y) ≤ f z :=
882880
@convex_on.le_on_segment _ _ _ _ (order_dual γ) _ _ _ f hf x y z hx hy hz
883881

882+
-- could be shown without contradiction but yeah
883+
lemma convex_on.le_left_of_right_le' (hf : convex_on s f) {x y : E} {a b : ℝ}
884+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 < a) (hb : 0 ≤ b) (hab : a + b = 1)
885+
(hxy : f y ≤ f (a • x + b • y)) :
886+
f (a • x + b • y) ≤ f x :=
887+
begin
888+
apply le_of_not_lt (λ h, lt_irrefl (f (a • x + b • y)) _),
889+
calc
890+
f (a • x + b • y)
891+
≤ a • f x + b • f y : hf.2 hx hy ha.le hb hab
892+
... < a • f (a • x + b • y) + b • f (a • x + b • y)
893+
: add_lt_add_of_lt_of_le (smul_lt_smul_of_pos h ha) (smul_le_smul_of_nonneg hxy hb)
894+
... = f (a • x + b • y) : by rw [←add_smul, hab, one_smul],
895+
end
896+
897+
lemma concave_on.left_le_of_le_right' (hf : concave_on s f) {x y : E} {a b : ℝ}
898+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 < a) (hb : 0 ≤ b) (hab : a + b = 1)
899+
(hxy : f (a • x + b • y) ≤ f y) :
900+
f x ≤ f (a • x + b • y) :=
901+
@convex_on.le_left_of_right_le' _ _ _ _ (order_dual γ) _ _ _ f hf x y a b hx hy ha hb hab hxy
902+
903+
lemma convex_on.le_right_of_left_le' (hf : convex_on s f) {x y : E} {a b : ℝ}
904+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 < b) (hab : a + b = 1)
905+
(hxy : f x ≤ f (a • x + b • y)) :
906+
f (a • x + b • y) ≤ f y :=
907+
begin
908+
rw add_comm at ⊢ hab hxy,
909+
exact hf.le_left_of_right_le' hy hx hb ha hab hxy,
910+
end
911+
912+
lemma concave_on.le_right_of_left_le' (hf : concave_on s f) {x y : E} {a b : ℝ}
913+
(hx : x ∈ s) (hy : y ∈ s) (ha : 0 ≤ a) (hb : 0 < b) (hab : a + b = 1)
914+
(hxy : f (a • x + b • y) ≤ f x) :
915+
f y ≤ f (a • x + b • y) :=
916+
@convex_on.le_right_of_left_le' _ _ _ _ (order_dual γ) _ _ _ f hf x y a b hx hy ha hb hab hxy
917+
918+
lemma convex_on.le_left_of_right_le (hf : convex_on s f) {x y z : E} (hx : x ∈ s)
919+
(hy : y ∈ s) (hz : z ∈ open_segment x y) (hyz : f y ≤ f z) :
920+
f z ≤ f x :=
921+
begin
922+
obtain ⟨a, b, ha, hb, hab, rfl⟩ := hz,
923+
exact hf.le_left_of_right_le' hx hy ha hb.le hab hyz,
924+
end
925+
926+
lemma concave_on.left_le_of_le_right (hf : concave_on s f) {x y z : E} (hx : x ∈ s)
927+
(hy : y ∈ s) (hz : z ∈ open_segment x y) (hyz : f z ≤ f y) :
928+
f x ≤ f z :=
929+
@convex_on.le_left_of_right_le _ _ _ _ (order_dual γ) _ _ _ f hf x y z hx hy hz hyz
930+
931+
lemma convex_on.le_right_of_left_le (hf : convex_on s f) {x y z : E} (hx : x ∈ s)
932+
(hy : y ∈ s) (hz : z ∈ open_segment x y) (hxz : f x ≤ f z) :
933+
f z ≤ f y :=
934+
begin
935+
obtain ⟨a, b, ha, hb, hab, rfl⟩ := hz,
936+
exact hf.le_right_of_left_le' hx hy ha.le hb hab hxz,
937+
end
938+
939+
lemma concave_on.le_right_of_left_le (hf : concave_on s f) {x y z : E} (hx : x ∈ s)
940+
(hy : y ∈ s) (hz : z ∈ open_segment x y) (hxz : f z ≤ f x) :
941+
f y ≤ f z :=
942+
@convex_on.le_right_of_left_le _ _ _ _ (order_dual γ) _ _ _ f hf x y z hx hy hz hxz
943+
944+
end linear_order
945+
884946
lemma convex_on.convex_le [ordered_module ℝ β] {f : E → β} (hf : convex_on s f) (r : β) :
885947
convex {x ∈ s | f x ≤ r} :=
886-
convex_iff_segment_subset.2 $ λ x y hx hy z hz,
948+
λ x y hx hy a b ha hb hab,
887949
begin
888-
refine ⟨hf.1.segment_subset hx.1 hy.1 hz,_⟩,
889-
rcases hz with ⟨za,zb,hza,hzb,hzazb,H⟩,
890-
rw ←H,
950+
refine ⟨hf.1 hx.1 hy.1 ha hb hab, _⟩,
891951
calc
892-
f (za • x + zb • y) ≤ za • (f x) + zb • (f y) : hf.2 hx.1 hy.1 hza hzb hzazb
893-
... ≤ za • r + zb • r : add_le_add (smul_le_smul_of_nonneg hx.2 hza)
894-
(smul_le_smul_of_nonneg hy.2 hzb)
895-
... ≤ r : by simp [←add_smul, hzazb]
952+
f (a • x + b • y) ≤ a • (f x) + b • (f y) : hf.2 hx.1 hy.1 ha hb hab
953+
... ≤ a • r + b • r : add_le_add (smul_le_smul_of_nonneg hx.2 ha)
954+
(smul_le_smul_of_nonneg hy.2 hb)
955+
... ≤ r : by simp [←add_smul, hab]
896956
end
897957

898958
lemma concave_on.concave_le [ordered_module ℝ β] {f : E → β} (hf : concave_on s f) (r : β) :
@@ -905,7 +965,6 @@ lemma convex_on.convex_lt {γ : Type*} [ordered_cancel_add_comm_monoid γ]
905965
begin
906966
intros a b as bs xa xb hxa hxb hxaxb,
907967
refine ⟨hf.1 as.1 bs.1 hxa hxb hxaxb, _⟩,
908-
dsimp,
909968
by_cases H : xa = 0,
910969
{ have H' : xb = 1 := by rwa [H, zero_add] at hxaxb,
911970
rw [H, H', zero_smul, one_smul, zero_add],

0 commit comments

Comments
 (0)