| 21776 |      1 | (*  Title       : FrechetDeriv.thy
 | 
|  |      2 |     ID          : $Id$
 | 
|  |      3 |     Author      : Brian Huffman
 | 
|  |      4 | *)
 | 
|  |      5 | 
 | 
|  |      6 | header {* Frechet Derivative *}
 | 
|  |      7 | 
 | 
|  |      8 | theory FrechetDeriv
 | 
|  |      9 | imports Lim
 | 
|  |     10 | begin
 | 
|  |     11 | 
 | 
|  |     12 | definition
 | 
|  |     13 |   fderiv ::
 | 
|  |     14 |   "['a::real_normed_vector \<Rightarrow> 'b::real_normed_vector, 'a, 'a \<Rightarrow> 'b] \<Rightarrow> bool"
 | 
|  |     15 |     -- {* Frechet derivative: D is derivative of function f at x *}
 | 
|  |     16 |           ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
 | 
|  |     17 |   "FDERIV f x :> D = (bounded_linear D \<and>
 | 
|  |     18 |     (\<lambda>h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)"
 | 
|  |     19 | 
 | 
|  |     20 | lemma FDERIV_I:
 | 
|  |     21 |   "\<lbrakk>bounded_linear D; (\<lambda>h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\<rbrakk>
 | 
|  |     22 |    \<Longrightarrow> FDERIV f x :> D"
 | 
|  |     23 | by (simp add: fderiv_def)
 | 
|  |     24 | 
 | 
|  |     25 | lemma FDERIV_D:
 | 
|  |     26 |   "FDERIV f x :> D \<Longrightarrow> (\<lambda>h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0"
 | 
|  |     27 | by (simp add: fderiv_def)
 | 
|  |     28 | 
 | 
|  |     29 | lemma FDERIV_bounded_linear: "FDERIV f x :> D \<Longrightarrow> bounded_linear D"
 | 
|  |     30 | by (simp add: fderiv_def)
 | 
|  |     31 | 
 | 
|  |     32 | lemma bounded_linear_zero:
 | 
|  |     33 |   "bounded_linear (\<lambda>x::'a::real_normed_vector. 0::'b::real_normed_vector)"
 | 
|  |     34 | proof (unfold_locales)
 | 
|  |     35 |   show "(0::'b) = 0 + 0" by simp
 | 
|  |     36 |   fix r show "(0::'b) = scaleR r 0" by simp
 | 
|  |     37 |   have "\<forall>x::'a. norm (0::'b) \<le> norm x * 0" by simp
 | 
|  |     38 |   thus "\<exists>K. \<forall>x::'a. norm (0::'b) \<le> norm x * K" ..
 | 
|  |     39 | qed
 | 
|  |     40 | 
 | 
|  |     41 | lemma FDERIV_const: "FDERIV (\<lambda>x. k) x :> (\<lambda>h. 0)"
 | 
|  |     42 | by (simp add: fderiv_def bounded_linear_zero)
 | 
|  |     43 | 
 | 
|  |     44 | lemma bounded_linear_ident:
 | 
|  |     45 |   "bounded_linear (\<lambda>x::'a::real_normed_vector. x)"
 | 
|  |     46 | proof (unfold_locales)
 | 
|  |     47 |   fix x y :: 'a show "x + y = x + y" by simp
 | 
|  |     48 |   fix r and x :: 'a show "scaleR r x = scaleR r x" by simp
 | 
|  |     49 |   have "\<forall>x::'a. norm x \<le> norm x * 1" by simp
 | 
|  |     50 |   thus "\<exists>K. \<forall>x::'a. norm x \<le> norm x * K" ..
 | 
|  |     51 | qed
 | 
|  |     52 | 
 | 
|  |     53 | lemma FDERIV_ident: "FDERIV (\<lambda>x. x) x :> (\<lambda>h. h)"
 | 
|  |     54 | by (simp add: fderiv_def bounded_linear_ident)
 | 
|  |     55 | 
 | 
|  |     56 | subsection {* Addition *}
 | 
|  |     57 | 
 | 
|  |     58 | lemma add_diff_add:
 | 
|  |     59 |   fixes a b c d :: "'a::ab_group_add"
 | 
|  |     60 |   shows "(a + c) - (b + d) = (a - b) + (c - d)"
 | 
|  |     61 | by simp
 | 
|  |     62 | 
 | 
|  |     63 | lemma bounded_linear_add:
 | 
|  |     64 |   includes bounded_linear f
 | 
|  |     65 |   includes bounded_linear g
 | 
|  |     66 |   shows "bounded_linear (\<lambda>x. f x + g x)"
 | 
|  |     67 | apply (unfold_locales)
 | 
|  |     68 | apply (simp only: f.add g.add add_ac)
 | 
|  |     69 | apply (simp only: f.scaleR g.scaleR scaleR_right_distrib)
 | 
| 22720 |     70 | apply (rule f.pos_bounded [THEN exE], rename_tac Kf)
 | 
|  |     71 | apply (rule g.pos_bounded [THEN exE], rename_tac Kg)
 | 
| 21776 |     72 | apply (rule_tac x="Kf + Kg" in exI, safe)
 | 
|  |     73 | apply (subst right_distrib)
 | 
|  |     74 | apply (rule order_trans [OF norm_triangle_ineq])
 | 
|  |     75 | apply (rule add_mono, erule spec, erule spec)
 | 
|  |     76 | done
 | 
|  |     77 | 
 | 
|  |     78 | lemma norm_ratio_ineq:
 | 
|  |     79 |   fixes x y :: "'a::real_normed_vector"
 | 
|  |     80 |   fixes h :: "'b::real_normed_vector"
 | 
|  |     81 |   shows "norm (x + y) / norm h \<le> norm x / norm h + norm y / norm h"
 | 
|  |     82 | apply (rule ord_le_eq_trans)
 | 
|  |     83 | apply (rule divide_right_mono)
 | 
|  |     84 | apply (rule norm_triangle_ineq)
 | 
|  |     85 | apply (rule norm_ge_zero)
 | 
|  |     86 | apply (rule add_divide_distrib)
 | 
|  |     87 | done
 | 
|  |     88 | 
 | 
|  |     89 | lemma FDERIV_add:
 | 
|  |     90 |   assumes f: "FDERIV f x :> F"
 | 
|  |     91 |   assumes g: "FDERIV g x :> G"
 | 
|  |     92 |   shows "FDERIV (\<lambda>x. f x + g x) x :> (\<lambda>h. F h + G h)"
 | 
|  |     93 | proof (rule FDERIV_I)
 | 
|  |     94 |   show "bounded_linear (\<lambda>h. F h + G h)"
 | 
|  |     95 |     apply (rule bounded_linear_add)
 | 
|  |     96 |     apply (rule FDERIV_bounded_linear [OF f])
 | 
|  |     97 |     apply (rule FDERIV_bounded_linear [OF g])
 | 
|  |     98 |     done
 | 
|  |     99 | next
 | 
|  |    100 |   have f': "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0"
 | 
|  |    101 |     using f by (rule FDERIV_D)
 | 
|  |    102 |   have g': "(\<lambda>h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0"
 | 
|  |    103 |     using g by (rule FDERIV_D)
 | 
|  |    104 |   from f' g'
 | 
|  |    105 |   have "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h
 | 
|  |    106 |            + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0"
 | 
|  |    107 |     by (rule LIM_add_zero)
 | 
|  |    108 |   thus "(\<lambda>h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h))
 | 
|  |    109 |            / norm h) -- 0 --> 0"
 | 
|  |    110 |     apply (rule real_LIM_sandwich_zero)
 | 
|  |    111 |      apply (simp add: divide_nonneg_pos)
 | 
|  |    112 |     apply (simp only: add_diff_add)
 | 
|  |    113 |     apply (rule norm_ratio_ineq)
 | 
|  |    114 |     done
 | 
|  |    115 | qed
 | 
|  |    116 | 
 | 
|  |    117 | subsection {* Subtraction *}
 | 
|  |    118 | 
 | 
|  |    119 | lemma bounded_linear_minus:
 | 
|  |    120 |   includes bounded_linear f
 | 
|  |    121 |   shows "bounded_linear (\<lambda>x. - f x)"
 | 
|  |    122 | apply (unfold_locales)
 | 
|  |    123 | apply (simp add: f.add)
 | 
|  |    124 | apply (simp add: f.scaleR)
 | 
|  |    125 | apply (simp add: f.bounded)
 | 
|  |    126 | done
 | 
|  |    127 | 
 | 
|  |    128 | lemma FDERIV_minus:
 | 
|  |    129 |   "FDERIV f x :> F \<Longrightarrow> FDERIV (\<lambda>x. - f x) x :> (\<lambda>h. - F h)"
 | 
|  |    130 | apply (rule FDERIV_I)
 | 
|  |    131 | apply (rule bounded_linear_minus)
 | 
|  |    132 | apply (erule FDERIV_bounded_linear)
 | 
|  |    133 | apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel)
 | 
|  |    134 | done
 | 
|  |    135 | 
 | 
|  |    136 | lemma FDERIV_diff:
 | 
|  |    137 |   "\<lbrakk>FDERIV f x :> F; FDERIV g x :> G\<rbrakk>
 | 
|  |    138 |    \<Longrightarrow> FDERIV (\<lambda>x. f x - g x) x :> (\<lambda>h. F h - G h)"
 | 
|  |    139 | by (simp only: diff_minus FDERIV_add FDERIV_minus)
 | 
|  |    140 | 
 | 
|  |    141 | subsection {* Continuity *}
 | 
|  |    142 | 
 | 
|  |    143 | lemma FDERIV_isCont:
 | 
|  |    144 |   assumes f: "FDERIV f x :> F"
 | 
|  |    145 |   shows "isCont f x"
 | 
|  |    146 | proof -
 | 
|  |    147 |   from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear)
 | 
|  |    148 |   have "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0"
 | 
|  |    149 |     by (rule FDERIV_D [OF f])
 | 
|  |    150 |   hence "(\<lambda>h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0"
 | 
|  |    151 |     by (intro LIM_mult_zero LIM_norm_zero LIM_self)
 | 
|  |    152 |   hence "(\<lambda>h. norm (f (x + h) - f x - F h)) -- 0 --> 0"
 | 
|  |    153 |     by (simp cong: LIM_cong)
 | 
|  |    154 |   hence "(\<lambda>h. f (x + h) - f x - F h) -- 0 --> 0"
 | 
|  |    155 |     by (rule LIM_norm_zero_cancel)
 | 
|  |    156 |   hence "(\<lambda>h. f (x + h) - f x - F h + F h) -- 0 --> 0"
 | 
|  |    157 |     by (intro LIM_add_zero F.LIM_zero LIM_self)
 | 
|  |    158 |   hence "(\<lambda>h. f (x + h) - f x) -- 0 --> 0"
 | 
|  |    159 |     by simp
 | 
|  |    160 |   thus "isCont f x"
 | 
|  |    161 |     unfolding isCont_iff by (rule LIM_zero_cancel)
 | 
|  |    162 | qed
 | 
|  |    163 | 
 | 
|  |    164 | subsection {* Composition *}
 | 
|  |    165 | 
 | 
|  |    166 | lemma real_divide_cancel_lemma:
 | 
|  |    167 |   fixes a b c :: real
 | 
|  |    168 |   shows "(b = 0 \<Longrightarrow> a = 0) \<Longrightarrow> (a / b) * (b / c) = a / c"
 | 
|  |    169 | by simp
 | 
|  |    170 | 
 | 
|  |    171 | lemma bounded_linear_compose:
 | 
|  |    172 |   includes bounded_linear f
 | 
|  |    173 |   includes bounded_linear g
 | 
|  |    174 |   shows "bounded_linear (\<lambda>x. f (g x))"
 | 
|  |    175 | proof (unfold_locales)
 | 
|  |    176 |   fix x y show "f (g (x + y)) = f (g x) + f (g y)"
 | 
|  |    177 |     by (simp only: f.add g.add)
 | 
|  |    178 | next
 | 
|  |    179 |   fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))"
 | 
|  |    180 |     by (simp only: f.scaleR g.scaleR)
 | 
|  |    181 | next
 | 
|  |    182 |   from f.pos_bounded
 | 
|  |    183 |   obtain Kf where f: "\<And>x. norm (f x) \<le> norm x * Kf" and Kf: "0 < Kf" by fast
 | 
|  |    184 |   from g.pos_bounded
 | 
|  |    185 |   obtain Kg where g: "\<And>x. norm (g x) \<le> norm x * Kg" by fast
 | 
|  |    186 |   show "\<exists>K. \<forall>x. norm (f (g x)) \<le> norm x * K"
 | 
|  |    187 |   proof (intro exI allI)
 | 
|  |    188 |     fix x
 | 
|  |    189 |     have "norm (f (g x)) \<le> norm (g x) * Kf"
 | 
|  |    190 |       using f .
 | 
|  |    191 |     also have "\<dots> \<le> (norm x * Kg) * Kf"
 | 
|  |    192 |       using g Kf [THEN order_less_imp_le] by (rule mult_right_mono)
 | 
|  |    193 |     also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)"
 | 
|  |    194 |       by (rule mult_assoc)
 | 
|  |    195 |     finally show "norm (f (g x)) \<le> norm x * (Kg * Kf)" .
 | 
|  |    196 |   qed
 | 
|  |    197 | qed
 | 
|  |    198 | 
 | 
|  |    199 | lemma FDERIV_compose:
 | 
|  |    200 |   fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
 | 
|  |    201 |   fixes g :: "'b::real_normed_vector \<Rightarrow> 'c::real_normed_vector"
 | 
|  |    202 |   assumes f: "FDERIV f x :> F"
 | 
|  |    203 |   assumes g: "FDERIV g (f x) :> G"
 | 
|  |    204 |   shows "FDERIV (\<lambda>x. g (f x)) x :> (\<lambda>h. G (F h))"
 | 
|  |    205 | proof (rule FDERIV_I)
 | 
|  |    206 |   from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f]
 | 
|  |    207 |   show "bounded_linear (\<lambda>h. G (F h))"
 | 
|  |    208 |     by (rule bounded_linear_compose)
 | 
|  |    209 | next
 | 
|  |    210 |   let ?Rf = "\<lambda>h. f (x + h) - f x - F h"
 | 
|  |    211 |   let ?Rg = "\<lambda>k. g (f x + k) - g (f x) - G k"
 | 
|  |    212 |   let ?k = "\<lambda>h. f (x + h) - f x"
 | 
|  |    213 |   let ?Nf = "\<lambda>h. norm (?Rf h) / norm h"
 | 
|  |    214 |   let ?Ng = "\<lambda>h. norm (?Rg (?k h)) / norm (?k h)"
 | 
|  |    215 |   from f interpret F: bounded_linear ["F"] by (rule FDERIV_bounded_linear)
 | 
|  |    216 |   from g interpret G: bounded_linear ["G"] by (rule FDERIV_bounded_linear)
 | 
|  |    217 |   from F.bounded obtain kF where kF: "\<And>x. norm (F x) \<le> norm x * kF" by fast
 | 
|  |    218 |   from G.bounded obtain kG where kG: "\<And>x. norm (G x) \<le> norm x * kG" by fast
 | 
|  |    219 | 
 | 
|  |    220 |   let ?fun2 = "\<lambda>h. ?Nf h * kG + ?Ng h * (?Nf h + kF)"
 | 
|  |    221 | 
 | 
|  |    222 |   show "(\<lambda>h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0"
 | 
|  |    223 |   proof (rule real_LIM_sandwich_zero)
 | 
|  |    224 |     have Nf: "?Nf -- 0 --> 0"
 | 
|  |    225 |       using FDERIV_D [OF f] .
 | 
|  |    226 | 
 | 
|  |    227 |     have Ng1: "isCont (\<lambda>k. norm (?Rg k) / norm k) 0"
 | 
|  |    228 |       by (simp add: isCont_def FDERIV_D [OF g])
 | 
|  |    229 |     have Ng2: "?k -- 0 --> 0"
 | 
|  |    230 |       apply (rule LIM_zero)
 | 
|  |    231 |       apply (fold isCont_iff)
 | 
|  |    232 |       apply (rule FDERIV_isCont [OF f])
 | 
|  |    233 |       done
 | 
|  |    234 |     have Ng: "?Ng -- 0 --> 0"
 | 
|  |    235 |       using isCont_LIM_compose [OF Ng1 Ng2] by simp
 | 
|  |    236 | 
 | 
|  |    237 |     have "(\<lambda>h. ?Nf h * kG + ?Ng h * (?Nf h + kF))
 | 
|  |    238 |            -- 0 --> 0 * kG + 0 * (0 + kF)"
 | 
|  |    239 |       by (intro LIM_add LIM_mult LIM_const Nf Ng)
 | 
|  |    240 |     thus "(\<lambda>h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0"
 | 
|  |    241 |       by simp
 | 
|  |    242 |   next
 | 
|  |    243 |     fix h::'a assume h: "h \<noteq> 0"
 | 
|  |    244 |     thus "0 \<le> norm (g (f (x + h)) - g (f x) - G (F h)) / norm h"
 | 
|  |    245 |       by (simp add: divide_nonneg_pos)
 | 
|  |    246 |   next
 | 
|  |    247 |     fix h::'a assume h: "h \<noteq> 0"
 | 
|  |    248 |     have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)"
 | 
|  |    249 |       by (simp add: G.diff)
 | 
|  |    250 |     hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h
 | 
|  |    251 |            = norm (G (?Rf h) + ?Rg (?k h)) / norm h"
 | 
|  |    252 |       by (rule arg_cong)
 | 
|  |    253 |     also have "\<dots> \<le> norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h"
 | 
|  |    254 |       by (rule norm_ratio_ineq)
 | 
|  |    255 |     also have "\<dots> \<le> ?Nf h * kG + ?Ng h * (?Nf h + kF)"
 | 
|  |    256 |     proof (rule add_mono)
 | 
|  |    257 |       show "norm (G (?Rf h)) / norm h \<le> ?Nf h * kG"
 | 
|  |    258 |         apply (rule ord_le_eq_trans)
 | 
|  |    259 |         apply (rule divide_right_mono [OF kG norm_ge_zero])
 | 
|  |    260 |         apply simp
 | 
|  |    261 |         done
 | 
|  |    262 |     next
 | 
|  |    263 |       have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)"
 | 
|  |    264 |         apply (rule real_divide_cancel_lemma [symmetric])
 | 
|  |    265 |         apply (simp add: G.zero)
 | 
|  |    266 |         done
 | 
|  |    267 |       also have "\<dots> \<le> ?Ng h * (?Nf h + kF)"
 | 
|  |    268 |       proof (rule mult_left_mono)
 | 
|  |    269 |         have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h"
 | 
|  |    270 |           by simp
 | 
|  |    271 |         also have "\<dots> \<le> ?Nf h + norm (F h) / norm h"
 | 
|  |    272 |           by (rule norm_ratio_ineq)
 | 
|  |    273 |         also have "\<dots> \<le> ?Nf h + kF"
 | 
|  |    274 |           apply (rule add_left_mono)
 | 
|  |    275 |           apply (subst pos_divide_le_eq, simp add: h)
 | 
|  |    276 |           apply (subst mult_commute)
 | 
|  |    277 |           apply (rule kF)
 | 
|  |    278 |           done
 | 
|  |    279 |         finally show "norm (?k h) / norm h \<le> ?Nf h + kF" .
 | 
|  |    280 |       next
 | 
|  |    281 |         show "0 \<le> ?Ng h"
 | 
|  |    282 |         apply (case_tac "f (x + h) - f x = 0", simp)
 | 
|  |    283 |         apply (rule divide_nonneg_pos [OF norm_ge_zero])
 | 
|  |    284 |         apply simp
 | 
|  |    285 |         done
 | 
|  |    286 |       qed
 | 
|  |    287 |       finally show "norm (?Rg (?k h)) / norm h \<le> ?Ng h * (?Nf h + kF)" .
 | 
|  |    288 |     qed
 | 
|  |    289 |     finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h
 | 
|  |    290 |         \<le> ?Nf h * kG + ?Ng h * (?Nf h + kF)" .
 | 
|  |    291 |   qed
 | 
|  |    292 | qed
 | 
|  |    293 | 
 | 
|  |    294 | subsection {* Product Rule *}
 | 
|  |    295 | 
 | 
|  |    296 | lemma (in bounded_bilinear) FDERIV_lemma:
 | 
|  |    297 |   "a' ** b' - a ** b - (a ** B + A ** b)
 | 
|  |    298 |    = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)"
 | 
|  |    299 | by (simp add: diff_left diff_right)
 | 
|  |    300 | 
 | 
|  |    301 | lemma (in bounded_bilinear) FDERIV:
 | 
|  |    302 |   fixes x :: "'d::real_normed_vector"
 | 
|  |    303 |   assumes f: "FDERIV f x :> F"
 | 
|  |    304 |   assumes g: "FDERIV g x :> G"
 | 
|  |    305 |   shows "FDERIV (\<lambda>x. f x ** g x) x :> (\<lambda>h. f x ** G h + F h ** g x)"
 | 
|  |    306 | proof (rule FDERIV_I)
 | 
|  |    307 |   show "bounded_linear (\<lambda>h. f x ** G h + F h ** g x)"
 | 
|  |    308 |     apply (rule bounded_linear_add)
 | 
|  |    309 |     apply (rule bounded_linear_compose [OF bounded_linear_right])
 | 
|  |    310 |     apply (rule FDERIV_bounded_linear [OF g])
 | 
|  |    311 |     apply (rule bounded_linear_compose [OF bounded_linear_left])
 | 
|  |    312 |     apply (rule FDERIV_bounded_linear [OF f])
 | 
|  |    313 |     done
 | 
|  |    314 | next
 | 
|  |    315 |   from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]]
 | 
|  |    316 |   obtain KF where norm_F: "\<And>x. norm (F x) \<le> norm x * KF" by fast
 | 
|  |    317 | 
 | 
|  |    318 |   from pos_bounded obtain K where K: "0 < K" and norm_prod:
 | 
|  |    319 |     "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" by fast
 | 
|  |    320 | 
 | 
|  |    321 |   let ?Rf = "\<lambda>h. f (x + h) - f x - F h"
 | 
|  |    322 |   let ?Rg = "\<lambda>h. g (x + h) - g x - G h"
 | 
|  |    323 | 
 | 
|  |    324 |   let ?fun1 = "\<lambda>h.
 | 
|  |    325 |         norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) /
 | 
|  |    326 |         norm h"
 | 
|  |    327 | 
 | 
|  |    328 |   let ?fun2 = "\<lambda>h.
 | 
|  |    329 |         norm (f x) * (norm (?Rg h) / norm h) * K +
 | 
|  |    330 |         norm (?Rf h) / norm h * norm (g (x + h)) * K +
 | 
|  |    331 |         KF * norm (g (x + h) - g x) * K"
 | 
|  |    332 | 
 | 
|  |    333 |   have "?fun1 -- 0 --> 0"
 | 
|  |    334 |   proof (rule real_LIM_sandwich_zero)
 | 
|  |    335 |     from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]]
 | 
|  |    336 |     have "?fun2 -- 0 -->
 | 
|  |    337 |           norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K"
 | 
|  |    338 |       by (intro LIM_add LIM_mult LIM_const LIM_norm LIM_zero FDERIV_D)
 | 
|  |    339 |     thus "?fun2 -- 0 --> 0"
 | 
|  |    340 |       by simp
 | 
|  |    341 |   next
 | 
|  |    342 |     fix h::'d assume "h \<noteq> 0"
 | 
|  |    343 |     thus "0 \<le> ?fun1 h"
 | 
|  |    344 |       by (simp add: divide_nonneg_pos)
 | 
|  |    345 |   next
 | 
|  |    346 |     fix h::'d assume "h \<noteq> 0"
 | 
|  |    347 |     have "?fun1 h \<le> (norm (f x) * norm (?Rg h) * K +
 | 
|  |    348 |          norm (?Rf h) * norm (g (x + h)) * K +
 | 
|  |    349 |          norm h * KF * norm (g (x + h) - g x) * K) / norm h"
 | 
|  |    350 |       by (intro
 | 
|  |    351 |         divide_right_mono mult_mono'
 | 
|  |    352 |         order_trans [OF norm_triangle_ineq add_mono]
 | 
|  |    353 |         order_trans [OF norm_prod mult_right_mono]
 | 
|  |    354 |         mult_nonneg_nonneg order_refl norm_ge_zero norm_F
 | 
|  |    355 |         K [THEN order_less_imp_le]
 | 
|  |    356 |       )
 | 
|  |    357 |     also have "\<dots> = ?fun2 h"
 | 
|  |    358 |       by (simp add: add_divide_distrib)
 | 
|  |    359 |     finally show "?fun1 h \<le> ?fun2 h" .
 | 
|  |    360 |   qed
 | 
|  |    361 |   thus "(\<lambda>h.
 | 
|  |    362 |     norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x))
 | 
|  |    363 |     / norm h) -- 0 --> 0"
 | 
|  |    364 |     by (simp only: FDERIV_lemma)
 | 
|  |    365 | qed
 | 
|  |    366 | 
 | 
|  |    367 | lemmas FDERIV_mult = bounded_bilinear_mult.FDERIV
 | 
|  |    368 | 
 | 
|  |    369 | lemmas FDERIV_scaleR = bounded_bilinear_scaleR.FDERIV
 | 
|  |    370 | 
 | 
|  |    371 | subsection {* Powers *}
 | 
|  |    372 | 
 | 
|  |    373 | lemma FDERIV_power_Suc:
 | 
|  |    374 |   fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}"
 | 
|  |    375 |   shows "FDERIV (\<lambda>x. x ^ Suc n) x :> (\<lambda>h. (of_nat n + 1) * x ^ n * h)"
 | 
|  |    376 |  apply (induct n)
 | 
|  |    377 |   apply (simp add: power_Suc FDERIV_ident)
 | 
|  |    378 |  apply (drule FDERIV_mult [OF FDERIV_ident])
 | 
|  |    379 |  apply (simp only: of_nat_Suc left_distrib mult_left_one)
 | 
|  |    380 |  apply (simp only: power_Suc right_distrib mult_ac)
 | 
|  |    381 | done
 | 
|  |    382 | 
 | 
|  |    383 | lemma FDERIV_power:
 | 
|  |    384 |   fixes x :: "'a::{real_normed_algebra,recpower,comm_ring_1}"
 | 
|  |    385 |   shows "FDERIV (\<lambda>x. x ^ n) x :> (\<lambda>h. of_nat n * x ^ (n - 1) * h)"
 | 
|  |    386 | by (cases n, simp add: FDERIV_const, simp add: FDERIV_power_Suc)
 | 
|  |    387 | 
 | 
|  |    388 | subsection {* Inverse *}
 | 
|  |    389 | 
 | 
|  |    390 | lemma inverse_diff_inverse:
 | 
|  |    391 |   "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk>
 | 
|  |    392 |    \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)"
 | 
|  |    393 | by (simp add: right_diff_distrib left_diff_distrib mult_assoc)
 | 
|  |    394 | 
 | 
|  |    395 | lemmas bounded_linear_mult_const =
 | 
|  |    396 |   bounded_bilinear_mult.bounded_linear_left [THEN bounded_linear_compose]
 | 
|  |    397 | 
 | 
|  |    398 | lemmas bounded_linear_const_mult =
 | 
|  |    399 |   bounded_bilinear_mult.bounded_linear_right [THEN bounded_linear_compose]
 | 
|  |    400 | 
 | 
|  |    401 | lemma FDERIV_inverse:
 | 
|  |    402 |   fixes x :: "'a::real_normed_div_algebra"
 | 
|  |    403 |   assumes x: "x \<noteq> 0"
 | 
|  |    404 |   shows "FDERIV inverse x :> (\<lambda>h. - (inverse x * h * inverse x))"
 | 
|  |    405 |         (is "FDERIV ?inv _ :> _")
 | 
|  |    406 | proof (rule FDERIV_I)
 | 
|  |    407 |   show "bounded_linear (\<lambda>h. - (?inv x * h * ?inv x))"
 | 
|  |    408 |     apply (rule bounded_linear_minus)
 | 
|  |    409 |     apply (rule bounded_linear_mult_const)
 | 
|  |    410 |     apply (rule bounded_linear_const_mult)
 | 
|  |    411 |     apply (rule bounded_linear_ident)
 | 
|  |    412 |     done
 | 
|  |    413 | next
 | 
|  |    414 |   show "(\<lambda>h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h)
 | 
|  |    415 |         -- 0 --> 0"
 | 
|  |    416 |   proof (rule LIM_equal2)
 | 
|  |    417 |     show "0 < norm x" using x by simp
 | 
|  |    418 |   next
 | 
|  |    419 |     fix h::'a
 | 
|  |    420 |     assume 1: "h \<noteq> 0"
 | 
|  |    421 |     assume "norm (h - 0) < norm x"
 | 
|  |    422 |     hence "h \<noteq> -x" by clarsimp
 | 
|  |    423 |     hence 2: "x + h \<noteq> 0"
 | 
|  |    424 |       apply (rule contrapos_nn)
 | 
|  |    425 |       apply (rule sym)
 | 
|  |    426 |       apply (erule equals_zero_I)
 | 
|  |    427 |       done
 | 
|  |    428 |     show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h
 | 
|  |    429 |           = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h"
 | 
|  |    430 |       apply (subst inverse_diff_inverse [OF 2 x])
 | 
|  |    431 |       apply (subst minus_diff_minus)
 | 
|  |    432 |       apply (subst norm_minus_cancel)
 | 
|  |    433 |       apply (simp add: left_diff_distrib)
 | 
|  |    434 |       done
 | 
|  |    435 |   next
 | 
|  |    436 |     show "(\<lambda>h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h)
 | 
|  |    437 |           -- 0 --> 0"
 | 
|  |    438 |     proof (rule real_LIM_sandwich_zero)
 | 
|  |    439 |       show "(\<lambda>h. norm (?inv (x + h) - ?inv x) * norm (?inv x))
 | 
|  |    440 |             -- 0 --> 0"
 | 
|  |    441 |         apply (rule LIM_mult_left_zero)
 | 
|  |    442 |         apply (rule LIM_norm_zero)
 | 
|  |    443 |         apply (rule LIM_zero)
 | 
|  |    444 |         apply (rule LIM_offset_zero)
 | 
|  |    445 |         apply (rule LIM_inverse)
 | 
|  |    446 |         apply (rule LIM_self)
 | 
|  |    447 |         apply (rule x)
 | 
|  |    448 |         done
 | 
|  |    449 |     next
 | 
|  |    450 |       fix h::'a assume h: "h \<noteq> 0"
 | 
|  |    451 |       show "0 \<le> norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h"
 | 
|  |    452 |         apply (rule divide_nonneg_pos)
 | 
|  |    453 |         apply (rule norm_ge_zero)
 | 
|  |    454 |         apply (simp add: h)
 | 
|  |    455 |         done
 | 
|  |    456 |     next
 | 
|  |    457 |       fix h::'a assume h: "h \<noteq> 0"
 | 
|  |    458 |       have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h
 | 
|  |    459 |             \<le> norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h"
 | 
|  |    460 |         apply (rule divide_right_mono [OF _ norm_ge_zero])
 | 
|  |    461 |         apply (rule order_trans [OF norm_mult_ineq])
 | 
|  |    462 |         apply (rule mult_right_mono [OF _ norm_ge_zero])
 | 
|  |    463 |         apply (rule norm_mult_ineq)
 | 
|  |    464 |         done
 | 
|  |    465 |       also have "\<dots> = norm (?inv (x + h) - ?inv x) * norm (?inv x)"
 | 
|  |    466 |         by simp
 | 
|  |    467 |       finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h
 | 
|  |    468 |             \<le> norm (?inv (x + h) - ?inv x) * norm (?inv x)" .   
 | 
|  |    469 |     qed
 | 
|  |    470 |   qed
 | 
|  |    471 | qed
 | 
|  |    472 | 
 | 
|  |    473 | subsection {* Alternate definition *}
 | 
|  |    474 | 
 | 
|  |    475 | lemma field_fderiv_def:
 | 
|  |    476 |   fixes x :: "'a::real_normed_field" shows
 | 
|  |    477 |   "FDERIV f x :> (\<lambda>h. h * D) = (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D"
 | 
|  |    478 |  apply (unfold fderiv_def)
 | 
|  |    479 |  apply (simp add: bounded_bilinear_mult.bounded_linear_left)
 | 
|  |    480 |  apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric])
 | 
|  |    481 |  apply (subst diff_divide_distrib)
 | 
|  |    482 |  apply (subst times_divide_eq_left [symmetric])
 | 
| 23398 |    483 |  apply (simp cong: LIM_cong)
 | 
| 21776 |    484 |  apply (simp add: LIM_norm_zero_iff LIM_zero_iff)
 | 
|  |    485 | done
 | 
|  |    486 | 
 | 
|  |    487 | end
 |