Documentation

Mathlib.Analysis.Calculus.LocalExtr.Basic

Local extrema of differentiable functions #

Main definitions #

In a real normed space E we define posTangentConeAt (s : Set E) (x : E). This would be the same as tangentConeAt ℝ≥0 s x if we had a theory of normed semifields. This set is used in the proof of Fermat's Theorem (see below), and can be used to formalize Lagrange multipliers and/or Karush–Kuhn–Tucker conditions.

Main statements #

For each theorem name listed below, we also prove similar theorems for min, extr (if applicable), and fderiv/deriv instead of HasFDerivAt/HasDerivAt.

Implementation notes #

For each mathematical fact we prove several versions of its formalization:

For the fderiv*/deriv* versions we omit the differentiability condition whenever it is possible due to the fact that fderiv and deriv are defined to be zero for non-differentiable functions.

References #

Tags #

local extremum, tangent cone, Fermat's Theorem

Positive tangent cone #

def posTangentConeAt {E : Type u} [NormedAddCommGroup E] [NormedSpace E] (s : Set E) (x : E) :
Set E

"Positive" tangent cone to s at x; the only difference from tangentConeAt is that we require c n → ∞ instead of ‖c n‖ → ∞. One can think about posTangentConeAt as tangentConeAt NNReal but we have no theory of normed semifields yet.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    theorem mem_posTangentConeAt_of_segment_subset {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {s : Set E} {x : E} {y : E} (h : segment x y s) :
    theorem mem_posTangentConeAt_of_segment_subset' {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {s : Set E} {x : E} {y : E} (h : segment x (x + y) s) :
    theorem posTangentConeAt_univ {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {a : E} :
    posTangentConeAt Set.univ a = Set.univ

    Fermat's Theorem (vector space) #

    theorem IsLocalMaxOn.hasFDerivWithinAt_nonpos {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } {s : Set E} (h : IsLocalMaxOn f s a) (hf : HasFDerivWithinAt f f' s a) {y : E} (hy : y posTangentConeAt s a) :
    f' y 0

    If f has a local max on s at a, f' is the derivative of f at a within s, and y belongs to the positive tangent cone of s at a, then f' y ≤ 0.

    theorem IsLocalMaxOn.fderivWithin_nonpos {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {s : Set E} (h : IsLocalMaxOn f s a) {y : E} (hy : y posTangentConeAt s a) :
    (fderivWithin f s a) y 0

    If f has a local max on s at a and y belongs to the positive tangent cone of s at a, then f' y ≤ 0.

    theorem IsLocalMaxOn.hasFDerivWithinAt_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } {s : Set E} (h : IsLocalMaxOn f s a) (hf : HasFDerivWithinAt f f' s a) {y : E} (hy : y posTangentConeAt s a) (hy' : -y posTangentConeAt s a) :
    f' y = 0

    If f has a local max on s at a, f' is a derivative of f at a within s, and both y and -y belong to the positive tangent cone of s at a, then f' y ≤ 0.

    theorem IsLocalMaxOn.fderivWithin_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {s : Set E} (h : IsLocalMaxOn f s a) {y : E} (hy : y posTangentConeAt s a) (hy' : -y posTangentConeAt s a) :
    (fderivWithin f s a) y = 0

    If f has a local max on s at a and both y and -y belong to the positive tangent cone of s at a, then f' y = 0.

    theorem IsLocalMinOn.hasFDerivWithinAt_nonneg {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } {s : Set E} (h : IsLocalMinOn f s a) (hf : HasFDerivWithinAt f f' s a) {y : E} (hy : y posTangentConeAt s a) :
    0 f' y

    If f has a local min on s at a, f' is the derivative of f at a within s, and y belongs to the positive tangent cone of s at a, then 0 ≤ f' y.

    theorem IsLocalMinOn.fderivWithin_nonneg {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {s : Set E} (h : IsLocalMinOn f s a) {y : E} (hy : y posTangentConeAt s a) :
    0 (fderivWithin f s a) y

    If f has a local min on s at a and y belongs to the positive tangent cone of s at a, then 0 ≤ f' y.

    theorem IsLocalMinOn.hasFDerivWithinAt_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } {s : Set E} (h : IsLocalMinOn f s a) (hf : HasFDerivWithinAt f f' s a) {y : E} (hy : y posTangentConeAt s a) (hy' : -y posTangentConeAt s a) :
    f' y = 0

    If f has a local max on s at a, f' is a derivative of f at a within s, and both y and -y belong to the positive tangent cone of s at a, then f' y ≤ 0.

    theorem IsLocalMinOn.fderivWithin_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {s : Set E} (h : IsLocalMinOn f s a) {y : E} (hy : y posTangentConeAt s a) (hy' : -y posTangentConeAt s a) :
    (fderivWithin f s a) y = 0

    If f has a local min on s at a and both y and -y belong to the positive tangent cone of s at a, then f' y = 0.

    theorem IsLocalMin.hasFDerivAt_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } (h : IsLocalMin f a) (hf : HasFDerivAt f f' a) :
    f' = 0

    Fermat's Theorem: the derivative of a function at a local minimum equals zero.

    theorem IsLocalMin.fderiv_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} (h : IsLocalMin f a) :
    fderiv f a = 0

    Fermat's Theorem: the derivative of a function at a local minimum equals zero.

    theorem IsLocalMax.hasFDerivAt_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } (h : IsLocalMax f a) (hf : HasFDerivAt f f' a) :
    f' = 0

    Fermat's Theorem: the derivative of a function at a local maximum equals zero.

    theorem IsLocalMax.fderiv_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} (h : IsLocalMax f a) :
    fderiv f a = 0

    Fermat's Theorem: the derivative of a function at a local maximum equals zero.

    theorem IsLocalExtr.hasFDerivAt_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} {f' : E →L[] } (h : IsLocalExtr f a) :
    HasFDerivAt f f' af' = 0

    Fermat's Theorem: the derivative of a function at a local extremum equals zero.

    theorem IsLocalExtr.fderiv_eq_zero {E : Type u} [NormedAddCommGroup E] [NormedSpace E] {f : E} {a : E} (h : IsLocalExtr f a) :
    fderiv f a = 0

    Fermat's Theorem: the derivative of a function at a local extremum equals zero.

    Fermat's Theorem #

    theorem IsLocalMin.hasDerivAt_eq_zero {f : } {f' : } {a : } (h : IsLocalMin f a) (hf : HasDerivAt f f' a) :
    f' = 0

    Fermat's Theorem: the derivative of a function at a local minimum equals zero.

    theorem IsLocalMin.deriv_eq_zero {f : } {a : } (h : IsLocalMin f a) :
    deriv f a = 0

    Fermat's Theorem: the derivative of a function at a local minimum equals zero.

    theorem IsLocalMax.hasDerivAt_eq_zero {f : } {f' : } {a : } (h : IsLocalMax f a) (hf : HasDerivAt f f' a) :
    f' = 0

    Fermat's Theorem: the derivative of a function at a local maximum equals zero.

    theorem IsLocalMax.deriv_eq_zero {f : } {a : } (h : IsLocalMax f a) :
    deriv f a = 0

    Fermat's Theorem: the derivative of a function at a local maximum equals zero.

    theorem IsLocalExtr.hasDerivAt_eq_zero {f : } {f' : } {a : } (h : IsLocalExtr f a) :
    HasDerivAt f f' af' = 0

    Fermat's Theorem: the derivative of a function at a local extremum equals zero.

    theorem IsLocalExtr.deriv_eq_zero {f : } {a : } (h : IsLocalExtr f a) :
    deriv f a = 0

    Fermat's Theorem: the derivative of a function at a local extremum equals zero.