Documentation

Mathlib.Analysis.Calculus.Deriv.MeanValue

Mean value theorem #

In this file we prove Cauchy's and Lagrange's mean value theorems, and deduce some corollaries.

Cauchy's mean value theorem says that for two functions f and g that are continuous on [a, b] and are differentiable on (a, b), there exists a point c ∈ (a, b) such that f' c / g' c = (f b - f a) / (g b - g a). We formulate this theorem with both sides multiplied by the denominators, see exists_ratio_hasDerivAt_eq_ratio_slope, in order to avoid auxiliary conditions like g' c ≠ 0.

Lagrange's mean value theorem, see exists_hasDerivAt_eq_slope, says that for a function f that is continuous on [a, b] and is differentiable on (a, b), there exists a point c ∈ (a, b) such that f' c = (f b - f a) / (b - a).

Lagrange's MVT implies that (f b - f a) / (b - a) > C provided that f' c > C for all c ∈ (a, b), see mul_sub_lt_image_sub_of_lt_deriv, and other theorems for > / / < / .

In case C = 0, we deduce that a function with a positive derivative is strictly monotone, see strictMonoOn_of_deriv_pos and nearby theorems for other types of monotonicity.

We also prove that a real function whose derivative tends to infinity from the right at a point is not differentiable on the right at that point, and similarly differentiability on the left.

Main results #

Functions [a, b] → ℝ. #

theorem exists_ratio_hasDerivAt_eq_ratio_slope (f f' : ) {a b : } (hab : a < b) (hfc : ContinuousOn f (Set.Icc a b)) (hff' : xSet.Ioo a b, HasDerivAt f (f' x) x) (g g' : ) (hgc : ContinuousOn g (Set.Icc a b)) (hgg' : xSet.Ioo a b, HasDerivAt g (g' x) x) :
cSet.Ioo a b, (g b - g a) * f' c = (f b - f a) * g' c

Cauchy's Mean Value Theorem, HasDerivAt version.

theorem exists_ratio_hasDerivAt_eq_ratio_slope' (f f' : ) {a b : } (hab : a < b) (g g' : ) {lfa lga lfb lgb : } (hff' : xSet.Ioo a b, HasDerivAt f (f' x) x) (hgg' : xSet.Ioo a b, HasDerivAt g (g' x) x) (hfa : Filter.Tendsto f (nhdsWithin a (Set.Ioi a)) (nhds lfa)) (hga : Filter.Tendsto g (nhdsWithin a (Set.Ioi a)) (nhds lga)) (hfb : Filter.Tendsto f (nhdsWithin b (Set.Iio b)) (nhds lfb)) (hgb : Filter.Tendsto g (nhdsWithin b (Set.Iio b)) (nhds lgb)) :
cSet.Ioo a b, (lgb - lga) * f' c = (lfb - lfa) * g' c

Cauchy's Mean Value Theorem, extended HasDerivAt version.

theorem exists_hasDerivAt_eq_slope (f f' : ) {a b : } (hab : a < b) (hfc : ContinuousOn f (Set.Icc a b)) (hff' : xSet.Ioo a b, HasDerivAt f (f' x) x) :
cSet.Ioo a b, f' c = (f b - f a) / (b - a)

Lagrange's Mean Value Theorem, HasDerivAt version

theorem exists_ratio_deriv_eq_ratio_slope (f : ) {a b : } (hab : a < b) (hfc : ContinuousOn f (Set.Icc a b)) (hfd : DifferentiableOn f (Set.Ioo a b)) (g : ) (hgc : ContinuousOn g (Set.Icc a b)) (hgd : DifferentiableOn g (Set.Ioo a b)) :
cSet.Ioo a b, (g b - g a) * deriv f c = (f b - f a) * deriv g c

Cauchy's Mean Value Theorem, deriv version.

theorem exists_ratio_deriv_eq_ratio_slope' (f : ) {a b : } (hab : a < b) (g : ) {lfa lga lfb lgb : } (hdf : DifferentiableOn f (Set.Ioo a b)) (hdg : DifferentiableOn g (Set.Ioo a b)) (hfa : Filter.Tendsto f (nhdsWithin a (Set.Ioi a)) (nhds lfa)) (hga : Filter.Tendsto g (nhdsWithin a (Set.Ioi a)) (nhds lga)) (hfb : Filter.Tendsto f (nhdsWithin b (Set.Iio b)) (nhds lfb)) (hgb : Filter.Tendsto g (nhdsWithin b (Set.Iio b)) (nhds lgb)) :
cSet.Ioo a b, (lgb - lga) * deriv f c = (lfb - lfa) * deriv g c

Cauchy's Mean Value Theorem, extended deriv version.

theorem exists_deriv_eq_slope (f : ) {a b : } (hab : a < b) (hfc : ContinuousOn f (Set.Icc a b)) (hfd : DifferentiableOn f (Set.Ioo a b)) :
cSet.Ioo a b, deriv f c = (f b - f a) / (b - a)

Lagrange's Mean Value Theorem, deriv version.

theorem exists_deriv_eq_slope' (f : ) {a b : } (hab : a < b) (hfc : ContinuousOn f (Set.Icc a b)) (hfd : DifferentiableOn f (Set.Ioo a b)) :
cSet.Ioo a b, deriv f c = slope f a b

Lagrange's Mean Value Theorem, deriv version.

A real function whose derivative tends to infinity from the right at a point is not differentiable on the right at that point.

A real function whose derivative tends to minus infinity from the right at a point is not differentiable on the right at that point

A real function whose derivative tends to minus infinity from the left at a point is not differentiable on the left at that point

A real function whose derivative tends to infinity from the left at a point is not differentiable on the left at that point

theorem Convex.mul_sub_lt_image_sub_of_lt_deriv {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) {C : } (hf'_gt : xinterior D, C < deriv f x) (x : ) :
x DyD, x < yC * (y - x) < f y - f x

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and C < f', then f grows faster than C * x on D, i.e., C * (y - x) < f y - f x whenever x, y ∈ D, x < y.

theorem mul_sub_lt_image_sub_of_lt_deriv {f : } (hf : Differentiable f) {C : } (hf'_gt : ∀ (x : ), C < deriv f x) x y : (hxy : x < y) :
C * (y - x) < f y - f x

Let f : ℝ → ℝ be a differentiable function. If C < f', then f grows faster than C * x, i.e., C * (y - x) < f y - f x whenever x < y.

theorem Convex.mul_sub_le_image_sub_of_le_deriv {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) {C : } (hf'_ge : xinterior D, C deriv f x) (x : ) :
x DyD, x yC * (y - x) f y - f x

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and C ≤ f', then f grows at least as fast as C * x on D, i.e., C * (y - x) ≤ f y - f x whenever x, y ∈ D, x ≤ y.

theorem mul_sub_le_image_sub_of_le_deriv {f : } (hf : Differentiable f) {C : } (hf'_ge : ∀ (x : ), C deriv f x) x y : (hxy : x y) :
C * (y - x) f y - f x

Let f : ℝ → ℝ be a differentiable function. If C ≤ f', then f grows at least as fast as C * x, i.e., C * (y - x) ≤ f y - f x whenever x ≤ y.

theorem Convex.image_sub_lt_mul_sub_of_deriv_lt {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) {C : } (lt_hf' : xinterior D, deriv f x < C) (x : ) (hx : x D) (y : ) (hy : y D) (hxy : x < y) :
f y - f x < C * (y - x)

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' < C, then f grows slower than C * x on D, i.e., f y - f x < C * (y - x) whenever x, y ∈ D, x < y.

theorem image_sub_lt_mul_sub_of_deriv_lt {f : } (hf : Differentiable f) {C : } (lt_hf' : ∀ (x : ), deriv f x < C) x y : (hxy : x < y) :
f y - f x < C * (y - x)

Let f : ℝ → ℝ be a differentiable function. If f' < C, then f grows slower than C * x on D, i.e., f y - f x < C * (y - x) whenever x < y.

theorem Convex.image_sub_le_mul_sub_of_deriv_le {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) {C : } (le_hf' : xinterior D, deriv f x C) (x : ) (hx : x D) (y : ) (hy : y D) (hxy : x y) :
f y - f x C * (y - x)

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' ≤ C, then f grows at most as fast as C * x on D, i.e., f y - f x ≤ C * (y - x) whenever x, y ∈ D, x ≤ y.

theorem image_sub_le_mul_sub_of_deriv_le {f : } (hf : Differentiable f) {C : } (le_hf' : ∀ (x : ), deriv f x C) x y : (hxy : x y) :
f y - f x C * (y - x)

Let f : ℝ → ℝ be a differentiable function. If f' ≤ C, then f grows at most as fast as C * x, i.e., f y - f x ≤ C * (y - x) whenever x ≤ y.

theorem strictMonoOn_of_deriv_pos {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : xinterior D, 0 < deriv f x) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is positive, then f is a strictly monotone function on D. Note that we don't require differentiability explicitly as it already implied by the derivative being strictly positive.

theorem strictMono_of_deriv_pos {f : } (hf' : ∀ (x : ), 0 < deriv f x) :

Let f : ℝ → ℝ be a differentiable function. If f' is positive, then f is a strictly monotone function. Note that we don't require differentiability explicitly as it already implied by the derivative being strictly positive.

theorem strictMonoOn_of_hasDerivWithinAt_pos {D : Set } (hD : Convex D) {f f' : } (hf : ContinuousOn f D) (hf' : xinterior D, HasDerivWithinAt f (f' x) (interior D) x) (hf'₀ : xinterior D, 0 < f' x) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is strictly positive, then f is a strictly monotone function on D.

theorem strictMono_of_hasDerivAt_pos {f f' : } (hf : ∀ (x : ), HasDerivAt f (f' x) x) (hf' : ∀ (x : ), 0 < f' x) :

Let f : ℝ → ℝ be a differentiable function. If f' is strictly positive, then f is a strictly monotone function.

theorem monotoneOn_of_deriv_nonneg {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) (hf'_nonneg : xinterior D, 0 deriv f x) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is nonnegative, then f is a monotone function on D.

theorem monotone_of_deriv_nonneg {f : } (hf : Differentiable f) (hf' : ∀ (x : ), 0 deriv f x) :

Let f : ℝ → ℝ be a differentiable function. If f' is nonnegative, then f is a monotone function.

theorem monotoneOn_of_hasDerivWithinAt_nonneg {D : Set } (hD : Convex D) {f f' : } (hf : ContinuousOn f D) (hf' : xinterior D, HasDerivWithinAt f (f' x) (interior D) x) (hf'₀ : xinterior D, 0 f' x) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is nonnegative, then f is a monotone function on D.

theorem monotone_of_hasDerivAt_nonneg {f f' : } (hf : ∀ (x : ), HasDerivAt f (f' x) x) (hf' : 0 f') :

Let f : ℝ → ℝ be a differentiable function. If f' is nonnegative, then f is a monotone function.

theorem strictAntiOn_of_deriv_neg {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : xinterior D, deriv f x < 0) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is negative, then f is a strictly antitone function on D.

theorem strictAnti_of_deriv_neg {f : } (hf' : ∀ (x : ), deriv f x < 0) :

Let f : ℝ → ℝ be a differentiable function. If f' is negative, then f is a strictly antitone function. Note that we don't require differentiability explicitly as it already implied by the derivative being strictly negative.

theorem strictAntiOn_of_hasDerivWithinAt_neg {D : Set } (hD : Convex D) {f f' : } (hf : ContinuousOn f D) (hf' : xinterior D, HasDerivWithinAt f (f' x) (interior D) x) (hf'₀ : xinterior D, f' x < 0) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is strictly positive, then f is a strictly monotone function on D.

theorem strictAnti_of_hasDerivAt_neg {f f' : } (hf : ∀ (x : ), HasDerivAt f (f' x) x) (hf' : ∀ (x : ), f' x < 0) :

Let f : ℝ → ℝ be a differentiable function. If f' is strictly positive, then f is a strictly monotone function.

theorem antitoneOn_of_deriv_nonpos {D : Set } (hD : Convex D) {f : } (hf : ContinuousOn f D) (hf' : DifferentiableOn f (interior D)) (hf'_nonpos : xinterior D, deriv f x 0) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is nonpositive, then f is an antitone function on D.

theorem antitone_of_deriv_nonpos {f : } (hf : Differentiable f) (hf' : ∀ (x : ), deriv f x 0) :

Let f : ℝ → ℝ be a differentiable function. If f' is nonpositive, then f is an antitone function.

theorem antitoneOn_of_hasDerivWithinAt_nonpos {D : Set } (hD : Convex D) {f f' : } (hf : ContinuousOn f D) (hf' : xinterior D, HasDerivWithinAt f (f' x) (interior D) x) (hf'₀ : xinterior D, f' x 0) :

Let f be a function continuous on a convex (or, equivalently, connected) subset D of the real line. If f is differentiable on the interior of D and f' is nonpositive, then f is an antitone function on D.

theorem antitone_of_hasDerivAt_nonpos {f f' : } (hf : ∀ (x : ), HasDerivAt f (f' x) x) (hf' : f' 0) :

Let f : ℝ → ℝ be a differentiable function. If f' is nonpositive, then f is an antitone function.

Functions f : E → ℝ #

theorem domain_mvt {E : Type u_1} [NormedAddCommGroup E] [NormedSpace E] {f : E} {s : Set E} {x y : E} {f' : EE →L[] } (hf : xs, HasFDerivWithinAt f (f' x) s x) (hs : Convex s) (xs : x s) (ys : y s) :
zsegment x y, f y - f x = (f' z) (y - x)

Lagrange's Mean Value Theorem, applied to convex domains.