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