@@ -11,7 +11,8 @@ From mathcomp Require Import normedtype landau forms poly.
1111(* This file provides a theory of differentiation. It includes the standard *)
1212(* rules of differentiation (differential of a sum, of a product, of *)
1313(* exponentiation, of the inverse, etc.) as well as standard theorems (the *)
14- (* Extreme Value Theorem, Rolle's theorem, the Mean Value Theorem). *)
14+ (* Extreme Value Theorem, Rolle's theorem, the Mean Value Theorem, Cauchy's *)
15+ (* mean value theorem, L'Hopital's rule). *)
1516(* *)
1617(* Parsable notations (in all of the following, f is not supposed to be *)
1718(* differentiable): *)
@@ -1602,6 +1603,237 @@ have [c cab D] := MVT altb fdrvbl fcont.
16021603by exists c => //; rewrite in_itv /= ltW (itvP cab).
16031604Qed .
16041605
1606+ Section Cauchy_MVT.
1607+ Context {R : realType}.
1608+ Variables (f df g dg : R -> R) (a b c : R).
1609+ Hypothesis ab : a < b.
1610+ Hypotheses (cf : {within `[a, b], continuous f})
1611+ (cg : {within `[a, b], continuous g}).
1612+ Hypotheses (fdf : forall x, x \in `]a, b[%R -> is_derive x 1 f (df x))
1613+ (gdg : forall x, x \in `]a, b[%R -> is_derive x 1 g (dg x)).
1614+ Hypotheses (dg0 : forall x, x \in `]a, b[%R -> dg x != 0).
1615+
1616+ Lemma differentiable_subr_neq0 : g b - g a != 0.
1617+ Proof .
1618+ have [r] := MVT ab gdg cg; rewrite in_itv/= => /andP[ar rb] dgg.
1619+ by rewrite dgg mulf_neq0 ?dg0 ?in_itv/= ?ar//; rewrite subr_eq0 gt_eqF.
1620+ Qed .
1621+
1622+ Lemma cauchy_MVT :
1623+ exists2 c, c \in `]a, b[%R & df c / dg c = (f b - f a) / (g b - g a).
1624+ Proof .
1625+ pose h x := f x - ((f b - f a) / (g b - g a)) * g x.
1626+ have hder x : x \in `]a, b[%R -> derivable h x 1.
1627+ move=> xab; apply: derivableB => /=.
1628+ exact: (@ex_derive _ _ _ _ _ _ _ (fdf xab)).
1629+ by apply: derivableM => //; exact: (@ex_derive _ _ _ _ _ _ _ (gdg xab)).
1630+ have ch : {within `[a, b], continuous h}.
1631+ rewrite continuous_subspace_in => x xab.
1632+ by apply: cvgB; [exact: cf|apply: cvgM; [exact: cvg_cst|exact: cg]].
1633+ have /(Rolle ab hder ch)[x xab derh] : h a = h b.
1634+ rewrite /h; apply/eqP; rewrite subr_eq eq_sym -addrA eq_sym addrC -subr_eq.
1635+ rewrite -mulrN -mulrDr -(addrC (g a)) -[X in _ * X]opprB mulrN -mulrA.
1636+ by rewrite mulVf ?differentiable_subr_neq0// mulr1 opprB.
1637+ pose dh x := df x - (f b - f a) / (g b - g a) * dg x.
1638+ have his_der y : y \in `]a, b[%R -> is_derive x 1 h (dh x).
1639+ by move=> yab; apply: is_deriveB; [exact: fdf|apply: is_deriveZ; exact: gdg].
1640+ exists x => //.
1641+ have := @derive_val _ R _ _ _ _ _ (his_der _ xab).
1642+ rewrite (@derive_val _ R _ _ _ _ _ derh) => /esym/eqP.
1643+ by rewrite subr_eq add0r => /eqP ->; rewrite -mulrA divff ?mulr1//; exact: dg0.
1644+ Qed .
1645+
1646+ End Cauchy_MVT.
1647+
1648+ Section lhopital.
1649+ Context {R : realType}.
1650+ Variables (f df g dg : R -> R) (a : R) (U : set R) (Ua : nbhs a U).
1651+ Hypotheses (fdf : forall x, x \in U -> is_derive x 1 f (df x))
1652+ (gdg : forall x, x \in U -> is_derive x 1 g (dg x)).
1653+ Hypotheses (fa0 : f a = 0) (ga0 : g a = 0)
1654+ (cdg : \forall x \near a^', dg x != 0).
1655+
1656+ Lemma lhopital_right (l : R) :
1657+ df x / dg x @[x --> a^'+] --> l -> f x / g x @[x --> a^'+] --> l.
1658+ Proof .
1659+ case: cdg => r/= r0 cdg'.
1660+ move: Ua; rewrite filter_of_nearI => -[D /= D0 aDU].
1661+ near a^'+ => b.
1662+ have abf x : x \in `]a, b[%R -> {within `[a, x], continuous f}.
1663+ rewrite in_itv/= => /andP[ax xb].
1664+ apply: derivable_within_continuous => y; rewrite in_itv/= => /andP[ay yx].
1665+ apply: ex_derive.
1666+ apply: fdf.
1667+ rewrite inE; apply: aDU => /=.
1668+ rewrite ler0_norm? subr_le0//.
1669+ rewrite opprD opprK addrC ltrBlDr (le_lt_trans yx)// (lt_trans xb)//.
1670+ near: b; apply: nbhs_right_lt.
1671+ by rewrite ltrDr.
1672+ have abg x : x \in `]a, b[%R -> {within `[a, x], continuous g}.
1673+ rewrite in_itv/= => /andP[ax xb].
1674+ apply: derivable_within_continuous => y; rewrite in_itv/= => /andP[ay yx].
1675+ apply: ex_derive.
1676+ apply: gdg.
1677+ rewrite inE; apply: aDU => /=.
1678+ rewrite ler0_norm? subr_le0//.
1679+ rewrite opprD opprK addrC ltrBlDr (le_lt_trans yx)// (lt_trans xb)//.
1680+ near: b; apply: nbhs_right_lt.
1681+ by rewrite ltrDr.
1682+ have fdf' y : y \in `]a, b[%R -> is_derive y 1 f (df y).
1683+ rewrite in_itv/= => /andP[ay yb]; apply: fdf.
1684+ rewrite inE; apply: aDU => /=.
1685+ rewrite ltr0_norm ?subr_lt0//.
1686+ rewrite opprD opprK addrC ltrBlDr (lt_le_trans yb)//.
1687+ near: b; apply: nbhs_right_le.
1688+ by rewrite ltrDr.
1689+ have gdg' y : y \in `]a, b[%R -> is_derive y 1 g (dg y).
1690+ rewrite in_itv/= => /andP[ay yb]; apply: gdg.
1691+ rewrite inE; apply: aDU => /=.
1692+ rewrite ltr0_norm ?subr_lt0//.
1693+ rewrite opprD opprK addrC ltrBlDr (lt_le_trans yb)//.
1694+ near: b; apply: nbhs_right_le.
1695+ by rewrite ltrDr.
1696+ have {}dg0 y : y \in `]a, b[%R -> dg y != 0.
1697+ rewrite in_itv/= => /andP[ay yb].
1698+ apply: (cdg' y) => /=; last by rewrite gt_eqF.
1699+ rewrite ltr0_norm; last by rewrite subr_lt0.
1700+ rewrite opprB ltrBlDl (lt_trans yb)//.
1701+ near: b; apply: nbhs_right_lt.
1702+ by rewrite ltrDl.
1703+ move=> fgal.
1704+ have L : \forall x \near a^'+,
1705+ exists2 c, c \in `]a, x[%R & df c / dg c = f x / g x.
1706+ near=> x.
1707+ have /andP[ax xb] : a < x < b by exact/andP.
1708+ have {}fdf' y : y \in `]a, x[%R -> is_derive y 1 f (df y).
1709+ rewrite !in_itv/= => /andP[ay yx].
1710+ by apply: fdf'; rewrite in_itv/= ay/= (lt_trans yx).
1711+ have {}gdg' y : y \in `]a, x[%R -> is_derive y 1 g (dg y).
1712+ rewrite !in_itv/= => /andP[ay yx].
1713+ by apply: gdg'; rewrite in_itv/= ay/= (lt_trans yx).
1714+ have {}dg0 y : y \in `]a, x[%R -> dg y != 0.
1715+ rewrite in_itv/= => /andP[ya yx].
1716+ by apply: dg0; rewrite in_itv/= ya/= (lt_trans yx).
1717+ have {}axf : {within `[a, x], continuous f}.
1718+ rewrite continuous_subspace_in => y; rewrite inE/= in_itv/= => /andP[ay yx].
1719+ by apply: abf; rewrite in_itv/= xb andbT.
1720+ have {}axg : {within `[a, x], continuous g}.
1721+ rewrite continuous_subspace_in => y; rewrite inE/= in_itv/= => /andP[ay yx].
1722+ by apply: abg; rewrite in_itv/= xb andbT.
1723+ have := @cauchy_MVT _ f df g dg _ _ ax axf axg fdf' gdg' dg0.
1724+ by rewrite fa0 ga0 2!subr0.
1725+ apply/cvgrPdist_le => /= e e0.
1726+ move/cvgrPdist_le : fgal.
1727+ move=> /(_ _ e0)[r'/= r'0 fagl].
1728+ case: L => d /= d0 L.
1729+ near=> t.
1730+ have /= := L t.
1731+ have atd : `|a - t| < d.
1732+ rewrite ltr0_norm; last by rewrite subr_lt0.
1733+ rewrite opprB ltrBlDl.
1734+ near: t; apply: nbhs_right_lt.
1735+ by rewrite ltrDl.
1736+ have at_ : a < t by [].
1737+ move=> /(_ atd)/(_ at_)[c]; rewrite in_itv/= => /andP[ac ct <-].
1738+ apply: fagl => //=.
1739+ rewrite ltr0_norm; last by rewrite subr_lt0.
1740+ rewrite opprB ltrBlDl (lt_trans ct)//.
1741+ near: t; apply: nbhs_right_lt.
1742+ by rewrite ltrDl.
1743+ Unshelve. all: by end_near. Qed .
1744+
1745+ Lemma lhopital_left (l : R) :
1746+ df x / dg x @[x --> a^'-] --> l -> f x / g x @[x --> a^'-] --> l.
1747+ Proof .
1748+ case: cdg => r/= r0 cdg'.
1749+ move: Ua; rewrite filter_of_nearI => -[D /= D0 aDU].
1750+ near a^'- => b.
1751+ have baf x : x \in `]b, a[%R -> {within `[x, a], continuous f}.
1752+ rewrite in_itv/= => /andP[bx xa].
1753+ apply: derivable_within_continuous => y; rewrite in_itv/= => /andP[xy ya].
1754+ apply: ex_derive.
1755+ apply: fdf.
1756+ rewrite inE; apply: aDU => /=.
1757+ rewrite ger0_norm ?subr_ge0//.
1758+ rewrite ltrBlDr -ltrBlDl (lt_le_trans _ xy)// (le_lt_trans _ bx)//.
1759+ near: b; apply: nbhs_left_ge.
1760+ by rewrite ltrBlDl ltrDr.
1761+ have bag x : x \in `]b, a[%R -> {within `[x, a], continuous g}.
1762+ rewrite in_itv/= => /andP[bx xa].
1763+ apply: derivable_within_continuous => y; rewrite in_itv/= => /andP[xy ya].
1764+ apply: ex_derive.
1765+ apply: gdg.
1766+ rewrite inE; apply: aDU => /=.
1767+ rewrite ger0_norm ?subr_ge0//.
1768+ rewrite ltrBlDr -ltrBlDl (lt_le_trans _ xy)// (lt_trans _ bx)//.
1769+ near: b; apply: nbhs_left_gt.
1770+ by rewrite ltrBlDl ltrDr.
1771+ have fdf' y : y \in `]b, a[%R -> is_derive y 1 f (df y).
1772+ rewrite in_itv/= => /andP[by_ ya]; apply: fdf.
1773+ rewrite inE.
1774+ apply: aDU => /=.
1775+ rewrite gtr0_norm ?subr_gt0//.
1776+ rewrite ltrBlDl -ltrBlDr (le_lt_trans _ by_)//.
1777+ near: b; apply: nbhs_left_ge.
1778+ by rewrite ltrBlDr ltrDl.
1779+ have gdg' y : y \in `]b, a[%R -> is_derive y 1 g (dg y).
1780+ rewrite in_itv/= => /andP[by_ ya]; apply: gdg.
1781+ rewrite inE; apply: aDU => /=.
1782+ rewrite gtr0_norm ?subr_gt0//.
1783+ rewrite ltrBlDl -ltrBlDr (le_lt_trans _ by_)//.
1784+ near: b; apply: nbhs_left_ge.
1785+ by rewrite ltrBlDr ltrDl.
1786+ have {}dg0 y : y \in `]b, a[%R -> dg y != 0.
1787+ rewrite in_itv/= => /andP[by_ ya].
1788+ apply: (cdg' y) => /=; last by rewrite lt_eqF.
1789+ rewrite gtr0_norm; last by rewrite subr_gt0.
1790+ rewrite ltrBlDr -ltrBlDl (lt_trans _ by_)//.
1791+ near: b; apply: nbhs_left_gt.
1792+ by rewrite ltrBlDl ltrDr.
1793+ move=> fgal.
1794+ have L : \forall x \near a^'-,
1795+ exists2 c, c \in `]x, a[%R & df c / dg c = f x / g x.
1796+ near=> x.
1797+ have /andP[bx xa] : b < x < a by exact/andP.
1798+ have {}fdf' y : y \in `]x, a[%R -> is_derive y 1 f (df y).
1799+ rewrite in_itv/= => /andP[xy ya].
1800+ by apply: fdf'; rewrite in_itv/= ya andbT (lt_trans bx).
1801+ have {}gdg' y : y \in `]x, a[%R -> is_derive y 1 g (dg y).
1802+ rewrite in_itv/= => /andP[xy ya].
1803+ by apply: gdg'; rewrite in_itv/= ya andbT (lt_trans _ xy).
1804+ have {}dg0 y : y \in `]x, a[%R -> dg y != 0.
1805+ rewrite in_itv/= => /andP[xy ya].
1806+ by apply: dg0; rewrite in_itv/= ya andbT (lt_trans bx).
1807+ have {}xaf : {within `[x, a], continuous f}.
1808+ rewrite continuous_subspace_in => y; rewrite inE/= in_itv/= => /andP[xy ya].
1809+ by apply: baf; rewrite in_itv/= bx.
1810+ have {}xag : {within `[x, a], continuous g}.
1811+ rewrite continuous_subspace_in => y; rewrite inE/= in_itv/= => /andP[xy ya].
1812+ by apply: bag; rewrite in_itv/= bx.
1813+ have := @cauchy_MVT _ f df g dg _ _ xa xaf xag fdf' gdg' dg0.
1814+ by rewrite fa0 ga0 !sub0r divrN mulNr opprK.
1815+ apply/cvgrPdist_le => /= e e0.
1816+ move/cvgrPdist_le : fgal.
1817+ move=> /(_ _ e0)[r'/= r'0 fagl].
1818+ case: L => d /= d0 L.
1819+ near=> t.
1820+ have /= := L t.
1821+ have atd : `|a - t| < d.
1822+ rewrite gtr0_norm; last by rewrite subr_gt0.
1823+ rewrite ltrBlDr -ltrBlDl.
1824+ near: t; apply: nbhs_left_gt.
1825+ by rewrite ltrBlDl ltrDr.
1826+ have ta : t < a by [].
1827+ move=> /(_ atd)/(_ ta)[c]; rewrite in_itv/= => /andP[tc ca <-].
1828+ apply: fagl => //=.
1829+ rewrite gtr0_norm; last by rewrite subr_gt0.
1830+ rewrite ltrBlDr -ltrBlDl (lt_trans _ tc)//.
1831+ near: t; apply: nbhs_left_gt.
1832+ by rewrite ltrBlDl ltrDr.
1833+ Unshelve. all: by end_near. Qed .
1834+
1835+ End lhopital.
1836+
16051837Lemma ler0_derive1_nincr (R : realType) (f : R -> R) (a b : R) :
16061838 (forall x, x \in `]a, b[%R -> derivable f x 1) ->
16071839 (forall x, x \in `]a, b[%R -> f^`() x <= 0) ->
0 commit comments