diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index ed20d75ad8529..5ccc195ea333f 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -918,7 +918,9 @@ lemma concave_on.smul [ordered_module ℝ β] {f : E → β} {c : ℝ} (hc : 0 @convex_on.smul _ _ _ _ (order_dual β) _ _ _ f c hc hf section linear_order -variables {γ : Type*} [linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ] +section monoid + +variables {γ : Type*} [linear_ordered_add_comm_monoid γ] [module ℝ γ] [ordered_module ℝ γ] {f : E → γ} /-- A convex function on a segment is upper-bounded by the max of its endpoints. -/ @@ -945,13 +947,16 @@ lemma convex_on.le_on_segment (hf : convex_on s f) {x y z : E} let ⟨a, b, ha, hb, hab, hz⟩ := hz in hz ▸ hf.le_on_segment' hx hy ha hb hab /-- A concave function on a segment is lower-bounded by the min of its endpoints. -/ -lemma concave_on.le_on_segment {γ : Type*} - [linear_ordered_add_comm_group γ] [module ℝ γ] [ordered_module ℝ γ] - {f : E → γ} (hf : concave_on s f) {x y z : E} +lemma concave_on.le_on_segment {f : E → γ} (hf : concave_on s f) {x y z : E} (hx : x ∈ s) (hy : y ∈ s) (hz : z ∈ [x, y]) : min (f x) (f y) ≤ f z := @convex_on.le_on_segment _ _ _ _ (order_dual γ) _ _ _ f hf x y z hx hy hz +end monoid + +variables {γ : Type*} [linear_ordered_cancel_add_comm_monoid γ] [module ℝ γ] [ordered_module ℝ γ] + {f : E → γ} + -- could be shown without contradiction but yeah lemma convex_on.le_left_of_right_le' (hf : convex_on s f) {x y : E} {a b : ℝ} (hx : x ∈ s) (hy : y ∈ s) (ha : 0 < a) (hb : 0 ≤ b) (hab : a + b = 1)