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 #
exists_ratio_hasDerivAt_eq_ratio_slope
andexists_ratio_deriv_eq_ratio_slope
: Cauchy's Mean Value Theorem.exists_hasDerivAt_eq_slope
andexists_deriv_eq_slope
: Lagrange's Mean Value Theorem.domain_mvt
: Lagrange's Mean Value Theorem, applied to a segment in a convex domain.Convex.image_sub_lt_mul_sub_of_deriv_lt
,Convex.mul_sub_lt_image_sub_of_lt_deriv
,Convex.image_sub_le_mul_sub_of_deriv_le
,Convex.mul_sub_le_image_sub_of_le_deriv
, if∀ x, C (</≤/>/≥) (f' x)
, thenC * (y - x) (</≤/>/≥) (f y - f x)
wheneverx < y
.monotoneOn_of_deriv_nonneg
,antitoneOn_of_deriv_nonpos
,strictMono_of_deriv_pos
,strictAnti_of_deriv_neg
: if the derivative of a function is non-negative/non-positive/positive/negative, then the function is monotone/antitone/strictly monotone/strictly monotonically decreasing.convexOn_of_deriv
,convexOn_of_deriv2_nonneg
: if the derivative of a function is increasing or its second derivative is nonnegative, then the original function is convex.
Functions [a, b] → ℝ
. #
Cauchy's Mean Value Theorem, HasDerivAt
version.
Cauchy's Mean Value Theorem, extended HasDerivAt
version.
Lagrange's Mean Value Theorem, HasDerivAt
version
Cauchy's Mean Value Theorem, deriv
version.
Cauchy's Mean Value Theorem, extended 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
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
.
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
.
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
.
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
.
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.
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.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is strictly positive, then
f
is a strictly monotone function.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is nonnegative, then
f
is a monotone function.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is nonnegative, then
f
is a monotone function.
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
.
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.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is strictly positive, then
f
is a strictly monotone function.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is nonpositive, then
f
is an antitone function.
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
.
Let f : ℝ → ℝ
be a differentiable function. If f'
is nonpositive, then f
is an antitone
function.
Functions f : E → ℝ
#
Lagrange's Mean Value Theorem, applied to convex domains.