diff --git a/Mathlib/Analysis/Convex/Function.lean b/Mathlib/Analysis/Convex/Function.lean index b181d3772efe8a..fdbb6368e40a79 100644 --- a/Mathlib/Analysis/Convex/Function.lean +++ b/Mathlib/Analysis/Convex/Function.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, FranΓ§ois Dupuis -/ import Mathlib.Analysis.Convex.Basic +import Mathlib.Order.Filter.Extr import Mathlib.Tactic.GCongr #align_import analysis.convex.function from "leanprover-community/mathlib"@"92ca63f0fb391a9ca5f22d2409a6080e786d99f7" @@ -1127,9 +1128,35 @@ theorem OrderIso.concaveOn_symm (f : Ξ± ≃o Ξ²) (hf : ConvexOn π•œ univ f) : end OrderIso -section -variable [LinearOrderedField π•œ] [LinearOrderedCancelAddCommMonoid Ξ²] [Module π•œ Ξ²] [OrderedSMul π•œ Ξ²] +section LinearOrderedField +variable [LinearOrderedField π•œ] + +section OrderedAddCommMonoid +variable [OrderedAddCommMonoid Ξ²] [AddCommMonoid E] [SMul π•œ E] [Module π•œ Ξ²] [OrderedSMul π•œ Ξ²] + {f : E β†’ Ξ²} {s : Set E} {x y : E} + +/-- A strictly convex function admits at most one global minimum. -/ +lemma StrictConvexOn.eq_of_isMinOn (hf : StrictConvexOn π•œ s f) (hfx : IsMinOn f s x) + (hfy : IsMinOn f s y) (hx : x ∈ s) (hy : y ∈ s) : x = y := by + by_contra hxy + let z := (2 : π•œ)⁻¹ β€’ x + (2 : π•œ)⁻¹ β€’ y + have hz : z ∈ s := hf.1 hx hy (by norm_num) (by norm_num) $ by norm_num + refine lt_irrefl (f z) ?_ + calc + f z < _ := hf.2 hx hy hxy (by norm_num) (by norm_num) $ by norm_num + _ ≀ (2 : π•œ)⁻¹ β€’ f z + (2 : π•œ)⁻¹ β€’ f z := by gcongr; exacts [hfx hz, hfy hz] + _ = f z := by rw [←_root_.add_smul]; norm_num + +/-- A strictly concave function admits at most one global maximum. -/ +lemma StrictConcaveOn.eq_of_isMaxOn (hf : StrictConcaveOn π•œ s f) (hfx : IsMaxOn f s x) + (hfy : IsMaxOn f s y) (hx : x ∈ s) (hy : y ∈ s) : x = y := + hf.dual.eq_of_isMinOn hfx hfy hx hy + +end OrderedAddCommMonoid + +section LinearOrderedCancelAddCommMonoid +variable [LinearOrderedCancelAddCommMonoid Ξ²] [Module π•œ Ξ²] [OrderedSMul π•œ Ξ²] {x y z : π•œ} {s : Set π•œ} {f : π•œ β†’ Ξ²} theorem ConvexOn.le_right_of_left_le'' (hf : ConvexOn π•œ s f) (hx : x ∈ s) (hz : z ∈ s) (hxy : x < y) @@ -1154,4 +1181,5 @@ theorem ConcaveOn.left_le_of_le_right'' (hf : ConcaveOn π•œ s f) (hx : x ∈ s) hf.dual.le_left_of_right_le'' hx hz hxy hyz h #align concave_on.left_le_of_le_right'' ConcaveOn.left_le_of_le_right'' -end +end LinearOrderedCancelAddCommMonoid +end LinearOrderedField