huffman@21776: (* Title : FrechetDeriv.thy huffman@21776: Author : Brian Huffman huffman@21776: *) huffman@21776: huffman@21776: header {* Frechet Derivative *} huffman@21776: huffman@21776: theory FrechetDeriv huffman@44282: imports Complex_Main huffman@21776: begin huffman@21776: huffman@21776: definition huffman@21776: fderiv :: huffman@21776: "['a::real_normed_vector \ 'b::real_normed_vector, 'a, 'a \ 'b] \ bool" huffman@21776: -- {* Frechet derivative: D is derivative of function f at x *} huffman@21776: ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where huffman@21776: "FDERIV f x :> D = (bounded_linear D \ huffman@21776: (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" huffman@21776: huffman@21776: lemma FDERIV_I: huffman@21776: "\bounded_linear D; (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0\ huffman@21776: \ FDERIV f x :> D" huffman@21776: by (simp add: fderiv_def) huffman@21776: huffman@21776: lemma FDERIV_D: huffman@21776: "FDERIV f x :> D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0" huffman@21776: by (simp add: fderiv_def) huffman@21776: huffman@21776: lemma FDERIV_bounded_linear: "FDERIV f x :> D \ bounded_linear D" huffman@21776: by (simp add: fderiv_def) huffman@21776: huffman@44127: lemma bounded_linear_zero: "bounded_linear (\x. 0)" huffman@44127: by (rule bounded_linear_intro [where K=0], simp_all) huffman@21776: huffman@21776: lemma FDERIV_const: "FDERIV (\x. k) x :> (\h. 0)" huffman@44568: by (simp add: fderiv_def bounded_linear_zero tendsto_const) huffman@21776: huffman@44127: lemma bounded_linear_ident: "bounded_linear (\x. x)" huffman@44127: by (rule bounded_linear_intro [where K=1], simp_all) huffman@21776: huffman@21776: lemma FDERIV_ident: "FDERIV (\x. x) x :> (\h. h)" huffman@44568: by (simp add: fderiv_def bounded_linear_ident tendsto_const) huffman@21776: huffman@21776: subsection {* Addition *} huffman@21776: huffman@21776: lemma bounded_linear_add: ballarin@27611: assumes "bounded_linear f" ballarin@27611: assumes "bounded_linear g" huffman@21776: shows "bounded_linear (\x. f x + g x)" ballarin@27611: proof - ballarin@29233: interpret f: bounded_linear f by fact ballarin@29233: interpret g: bounded_linear g by fact ballarin@27611: show ?thesis apply (unfold_locales) ballarin@27611: apply (simp only: f.add g.add add_ac) ballarin@27611: apply (simp only: f.scaleR g.scaleR scaleR_right_distrib) ballarin@27611: apply (rule f.pos_bounded [THEN exE], rename_tac Kf) ballarin@27611: apply (rule g.pos_bounded [THEN exE], rename_tac Kg) ballarin@27611: apply (rule_tac x="Kf + Kg" in exI, safe) ballarin@27611: apply (subst right_distrib) ballarin@27611: apply (rule order_trans [OF norm_triangle_ineq]) ballarin@27611: apply (rule add_mono, erule spec, erule spec) ballarin@27611: done ballarin@27611: qed huffman@21776: huffman@21776: lemma norm_ratio_ineq: huffman@21776: fixes x y :: "'a::real_normed_vector" huffman@21776: fixes h :: "'b::real_normed_vector" huffman@21776: shows "norm (x + y) / norm h \ norm x / norm h + norm y / norm h" huffman@21776: apply (rule ord_le_eq_trans) huffman@21776: apply (rule divide_right_mono) huffman@21776: apply (rule norm_triangle_ineq) huffman@21776: apply (rule norm_ge_zero) huffman@21776: apply (rule add_divide_distrib) huffman@21776: done huffman@21776: huffman@21776: lemma FDERIV_add: huffman@21776: assumes f: "FDERIV f x :> F" huffman@21776: assumes g: "FDERIV g x :> G" huffman@21776: shows "FDERIV (\x. f x + g x) x :> (\h. F h + G h)" huffman@21776: proof (rule FDERIV_I) huffman@21776: show "bounded_linear (\h. F h + G h)" huffman@21776: apply (rule bounded_linear_add) huffman@21776: apply (rule FDERIV_bounded_linear [OF f]) huffman@21776: apply (rule FDERIV_bounded_linear [OF g]) huffman@21776: done huffman@21776: next huffman@21776: have f': "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" huffman@21776: using f by (rule FDERIV_D) huffman@21776: have g': "(\h. norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" huffman@21776: using g by (rule FDERIV_D) huffman@21776: from f' g' huffman@21776: have "(\h. norm (f (x + h) - f x - F h) / norm h huffman@21776: + norm (g (x + h) - g x - G h) / norm h) -- 0 --> 0" huffman@44568: by (rule tendsto_add_zero) huffman@21776: thus "(\h. norm (f (x + h) + g (x + h) - (f x + g x) - (F h + G h)) huffman@21776: / norm h) -- 0 --> 0" huffman@21776: apply (rule real_LIM_sandwich_zero) huffman@21776: apply (simp add: divide_nonneg_pos) huffman@21776: apply (simp only: add_diff_add) huffman@21776: apply (rule norm_ratio_ineq) huffman@21776: done huffman@21776: qed huffman@21776: huffman@21776: subsection {* Subtraction *} huffman@21776: huffman@21776: lemma bounded_linear_minus: ballarin@27611: assumes "bounded_linear f" huffman@21776: shows "bounded_linear (\x. - f x)" ballarin@27611: proof - ballarin@29233: interpret f: bounded_linear f by fact ballarin@27611: show ?thesis apply (unfold_locales) ballarin@27611: apply (simp add: f.add) ballarin@27611: apply (simp add: f.scaleR) ballarin@27611: apply (simp add: f.bounded) ballarin@27611: done ballarin@27611: qed huffman@21776: huffman@21776: lemma FDERIV_minus: huffman@21776: "FDERIV f x :> F \ FDERIV (\x. - f x) x :> (\h. - F h)" huffman@21776: apply (rule FDERIV_I) huffman@21776: apply (rule bounded_linear_minus) huffman@21776: apply (erule FDERIV_bounded_linear) huffman@21776: apply (simp only: fderiv_def minus_diff_minus norm_minus_cancel) huffman@21776: done huffman@21776: huffman@21776: lemma FDERIV_diff: huffman@21776: "\FDERIV f x :> F; FDERIV g x :> G\ huffman@21776: \ FDERIV (\x. f x - g x) x :> (\h. F h - G h)" huffman@21776: by (simp only: diff_minus FDERIV_add FDERIV_minus) huffman@21776: huffman@37729: subsection {* Uniqueness *} huffman@37729: huffman@37729: lemma FDERIV_zero_unique: huffman@37729: assumes "FDERIV (\x. 0) x :> F" shows "F = (\h. 0)" huffman@37729: proof - huffman@37729: interpret F: bounded_linear F huffman@37729: using assms by (rule FDERIV_bounded_linear) huffman@37729: let ?r = "\h. norm (F h) / norm h" huffman@37729: have *: "?r -- 0 --> 0" huffman@37729: using assms unfolding fderiv_def by simp huffman@37729: show "F = (\h. 0)" huffman@37729: proof huffman@37729: fix h show "F h = 0" huffman@37729: proof (rule ccontr) huffman@37729: assume "F h \ 0" huffman@37729: moreover from this have h: "h \ 0" huffman@37729: by (clarsimp simp add: F.zero) huffman@37729: ultimately have "0 < ?r h" huffman@37729: by (simp add: divide_pos_pos) huffman@37729: from LIM_D [OF * this] obtain s where s: "0 < s" huffman@37729: and r: "\x. x \ 0 \ norm x < s \ ?r x < ?r h" by auto huffman@37729: from dense [OF s] obtain t where t: "0 < t \ t < s" .. huffman@37729: let ?x = "scaleR (t / norm h) h" huffman@37729: have "?x \ 0" and "norm ?x < s" using t h by simp_all huffman@37729: hence "?r ?x < ?r h" by (rule r) huffman@37729: thus "False" using t h by (simp add: F.scaleR) huffman@37729: qed huffman@37729: qed huffman@37729: qed huffman@37729: huffman@37729: lemma FDERIV_unique: huffman@37729: assumes "FDERIV f x :> F" and "FDERIV f x :> F'" shows "F = F'" huffman@37729: proof - huffman@37729: have "FDERIV (\x. 0) x :> (\h. F h - F' h)" huffman@37729: using FDERIV_diff [OF assms] by simp huffman@37729: hence "(\h. F h - F' h) = (\h. 0)" huffman@37729: by (rule FDERIV_zero_unique) huffman@37729: thus "F = F'" nipkow@39302: unfolding fun_eq_iff right_minus_eq . huffman@37729: qed huffman@37729: huffman@21776: subsection {* Continuity *} huffman@21776: huffman@21776: lemma FDERIV_isCont: huffman@21776: assumes f: "FDERIV f x :> F" huffman@21776: shows "isCont f x" huffman@21776: proof - ballarin@29233: from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) huffman@21776: have "(\h. norm (f (x + h) - f x - F h) / norm h) -- 0 --> 0" huffman@21776: by (rule FDERIV_D [OF f]) huffman@21776: hence "(\h. norm (f (x + h) - f x - F h) / norm h * norm h) -- 0 --> 0" huffman@44568: by (intro tendsto_mult_zero tendsto_norm_zero tendsto_ident_at) huffman@21776: hence "(\h. norm (f (x + h) - f x - F h)) -- 0 --> 0" huffman@21776: by (simp cong: LIM_cong) huffman@21776: hence "(\h. f (x + h) - f x - F h) -- 0 --> 0" huffman@44568: by (rule tendsto_norm_zero_cancel) huffman@21776: hence "(\h. f (x + h) - f x - F h + F h) -- 0 --> 0" huffman@44568: by (intro tendsto_add_zero F.tendsto_zero tendsto_ident_at) huffman@21776: hence "(\h. f (x + h) - f x) -- 0 --> 0" huffman@21776: by simp huffman@21776: thus "isCont f x" huffman@21776: unfolding isCont_iff by (rule LIM_zero_cancel) huffman@21776: qed huffman@21776: huffman@21776: subsection {* Composition *} huffman@21776: huffman@21776: lemma real_divide_cancel_lemma: huffman@21776: fixes a b c :: real huffman@21776: shows "(b = 0 \ a = 0) \ (a / b) * (b / c) = a / c" huffman@21776: by simp huffman@21776: huffman@21776: lemma bounded_linear_compose: ballarin@27611: assumes "bounded_linear f" ballarin@27611: assumes "bounded_linear g" huffman@21776: shows "bounded_linear (\x. f (g x))" ballarin@27611: proof - ballarin@29233: interpret f: bounded_linear f by fact ballarin@29233: interpret g: bounded_linear g by fact ballarin@27611: show ?thesis proof (unfold_locales) ballarin@27611: fix x y show "f (g (x + y)) = f (g x) + f (g y)" ballarin@27611: by (simp only: f.add g.add) ballarin@27611: next ballarin@27611: fix r x show "f (g (scaleR r x)) = scaleR r (f (g x))" ballarin@27611: by (simp only: f.scaleR g.scaleR) ballarin@27611: next ballarin@27611: from f.pos_bounded ballarin@27611: obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by fast ballarin@27611: from g.pos_bounded ballarin@27611: obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by fast ballarin@27611: show "\K. \x. norm (f (g x)) \ norm x * K" ballarin@27611: proof (intro exI allI) ballarin@27611: fix x ballarin@27611: have "norm (f (g x)) \ norm (g x) * Kf" wenzelm@32960: using f . ballarin@27611: also have "\ \ (norm x * Kg) * Kf" wenzelm@32960: using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) ballarin@27611: also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" wenzelm@32960: by (rule mult_assoc) ballarin@27611: finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . ballarin@27611: qed huffman@21776: qed huffman@21776: qed huffman@21776: huffman@21776: lemma FDERIV_compose: huffman@21776: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" huffman@21776: fixes g :: "'b::real_normed_vector \ 'c::real_normed_vector" huffman@21776: assumes f: "FDERIV f x :> F" huffman@21776: assumes g: "FDERIV g (f x) :> G" huffman@21776: shows "FDERIV (\x. g (f x)) x :> (\h. G (F h))" huffman@21776: proof (rule FDERIV_I) huffman@21776: from FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f] huffman@21776: show "bounded_linear (\h. G (F h))" huffman@21776: by (rule bounded_linear_compose) huffman@21776: next huffman@21776: let ?Rf = "\h. f (x + h) - f x - F h" huffman@21776: let ?Rg = "\k. g (f x + k) - g (f x) - G k" huffman@21776: let ?k = "\h. f (x + h) - f x" huffman@21776: let ?Nf = "\h. norm (?Rf h) / norm h" huffman@21776: let ?Ng = "\h. norm (?Rg (?k h)) / norm (?k h)" wenzelm@30729: from f interpret F: bounded_linear "F" by (rule FDERIV_bounded_linear) wenzelm@30729: from g interpret G: bounded_linear "G" by (rule FDERIV_bounded_linear) huffman@21776: from F.bounded obtain kF where kF: "\x. norm (F x) \ norm x * kF" by fast huffman@21776: from G.bounded obtain kG where kG: "\x. norm (G x) \ norm x * kG" by fast huffman@21776: huffman@21776: let ?fun2 = "\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)" huffman@21776: huffman@21776: show "(\h. norm (g (f (x + h)) - g (f x) - G (F h)) / norm h) -- 0 --> 0" huffman@21776: proof (rule real_LIM_sandwich_zero) huffman@21776: have Nf: "?Nf -- 0 --> 0" huffman@21776: using FDERIV_D [OF f] . huffman@21776: huffman@21776: have Ng1: "isCont (\k. norm (?Rg k) / norm k) 0" huffman@21776: by (simp add: isCont_def FDERIV_D [OF g]) huffman@21776: have Ng2: "?k -- 0 --> 0" huffman@21776: apply (rule LIM_zero) huffman@21776: apply (fold isCont_iff) huffman@21776: apply (rule FDERIV_isCont [OF f]) huffman@21776: done huffman@21776: have Ng: "?Ng -- 0 --> 0" huffman@44568: using isCont_tendsto_compose [OF Ng1 Ng2] by simp huffman@21776: huffman@21776: have "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) huffman@21776: -- 0 --> 0 * kG + 0 * (0 + kF)" huffman@44568: by (intro tendsto_intros Nf Ng) huffman@21776: thus "(\h. ?Nf h * kG + ?Ng h * (?Nf h + kF)) -- 0 --> 0" huffman@21776: by simp huffman@21776: next huffman@21776: fix h::'a assume h: "h \ 0" huffman@21776: thus "0 \ norm (g (f (x + h)) - g (f x) - G (F h)) / norm h" huffman@21776: by (simp add: divide_nonneg_pos) huffman@21776: next huffman@21776: fix h::'a assume h: "h \ 0" huffman@21776: have "g (f (x + h)) - g (f x) - G (F h) = G (?Rf h) + ?Rg (?k h)" huffman@21776: by (simp add: G.diff) huffman@21776: hence "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h huffman@21776: = norm (G (?Rf h) + ?Rg (?k h)) / norm h" huffman@21776: by (rule arg_cong) huffman@21776: also have "\ \ norm (G (?Rf h)) / norm h + norm (?Rg (?k h)) / norm h" huffman@21776: by (rule norm_ratio_ineq) huffman@21776: also have "\ \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" huffman@21776: proof (rule add_mono) huffman@21776: show "norm (G (?Rf h)) / norm h \ ?Nf h * kG" huffman@21776: apply (rule ord_le_eq_trans) huffman@21776: apply (rule divide_right_mono [OF kG norm_ge_zero]) huffman@21776: apply simp huffman@21776: done huffman@21776: next huffman@21776: have "norm (?Rg (?k h)) / norm h = ?Ng h * (norm (?k h) / norm h)" huffman@21776: apply (rule real_divide_cancel_lemma [symmetric]) huffman@21776: apply (simp add: G.zero) huffman@21776: done huffman@21776: also have "\ \ ?Ng h * (?Nf h + kF)" huffman@21776: proof (rule mult_left_mono) huffman@21776: have "norm (?k h) / norm h = norm (?Rf h + F h) / norm h" huffman@21776: by simp huffman@21776: also have "\ \ ?Nf h + norm (F h) / norm h" huffman@21776: by (rule norm_ratio_ineq) huffman@21776: also have "\ \ ?Nf h + kF" huffman@21776: apply (rule add_left_mono) huffman@21776: apply (subst pos_divide_le_eq, simp add: h) huffman@21776: apply (subst mult_commute) huffman@21776: apply (rule kF) huffman@21776: done huffman@21776: finally show "norm (?k h) / norm h \ ?Nf h + kF" . huffman@21776: next huffman@21776: show "0 \ ?Ng h" huffman@21776: apply (case_tac "f (x + h) - f x = 0", simp) huffman@21776: apply (rule divide_nonneg_pos [OF norm_ge_zero]) huffman@21776: apply simp huffman@21776: done huffman@21776: qed huffman@21776: finally show "norm (?Rg (?k h)) / norm h \ ?Ng h * (?Nf h + kF)" . huffman@21776: qed huffman@21776: finally show "norm (g (f (x + h)) - g (f x) - G (F h)) / norm h huffman@21776: \ ?Nf h * kG + ?Ng h * (?Nf h + kF)" . huffman@21776: qed huffman@21776: qed huffman@21776: huffman@21776: subsection {* Product Rule *} huffman@21776: huffman@21776: lemma (in bounded_bilinear) FDERIV_lemma: huffman@21776: "a' ** b' - a ** b - (a ** B + A ** b) huffman@21776: = a ** (b' - b - B) + (a' - a - A) ** b' + A ** (b' - b)" huffman@21776: by (simp add: diff_left diff_right) huffman@21776: huffman@21776: lemma (in bounded_bilinear) FDERIV: huffman@21776: fixes x :: "'d::real_normed_vector" huffman@21776: assumes f: "FDERIV f x :> F" huffman@21776: assumes g: "FDERIV g x :> G" huffman@21776: shows "FDERIV (\x. f x ** g x) x :> (\h. f x ** G h + F h ** g x)" huffman@21776: proof (rule FDERIV_I) huffman@21776: show "bounded_linear (\h. f x ** G h + F h ** g x)" huffman@21776: apply (rule bounded_linear_add) huffman@21776: apply (rule bounded_linear_compose [OF bounded_linear_right]) huffman@21776: apply (rule FDERIV_bounded_linear [OF g]) huffman@21776: apply (rule bounded_linear_compose [OF bounded_linear_left]) huffman@21776: apply (rule FDERIV_bounded_linear [OF f]) huffman@21776: done huffman@21776: next huffman@21776: from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] huffman@21776: obtain KF where norm_F: "\x. norm (F x) \ norm x * KF" by fast huffman@21776: huffman@21776: from pos_bounded obtain K where K: "0 < K" and norm_prod: huffman@21776: "\a b. norm (a ** b) \ norm a * norm b * K" by fast huffman@21776: huffman@21776: let ?Rf = "\h. f (x + h) - f x - F h" huffman@21776: let ?Rg = "\h. g (x + h) - g x - G h" huffman@21776: huffman@21776: let ?fun1 = "\h. huffman@21776: norm (f x ** ?Rg h + ?Rf h ** g (x + h) + F h ** (g (x + h) - g x)) / huffman@21776: norm h" huffman@21776: huffman@21776: let ?fun2 = "\h. huffman@21776: norm (f x) * (norm (?Rg h) / norm h) * K + huffman@21776: norm (?Rf h) / norm h * norm (g (x + h)) * K + huffman@21776: KF * norm (g (x + h) - g x) * K" huffman@21776: huffman@21776: have "?fun1 -- 0 --> 0" huffman@21776: proof (rule real_LIM_sandwich_zero) huffman@21776: from f g isCont_iff [THEN iffD1, OF FDERIV_isCont [OF g]] huffman@21776: have "?fun2 -- 0 --> huffman@21776: norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K" huffman@44568: by (intro tendsto_intros LIM_zero FDERIV_D) huffman@21776: thus "?fun2 -- 0 --> 0" huffman@21776: by simp huffman@21776: next huffman@21776: fix h::'d assume "h \ 0" huffman@21776: thus "0 \ ?fun1 h" huffman@21776: by (simp add: divide_nonneg_pos) huffman@21776: next huffman@21776: fix h::'d assume "h \ 0" huffman@21776: have "?fun1 h \ (norm (f x) * norm (?Rg h) * K + huffman@21776: norm (?Rf h) * norm (g (x + h)) * K + huffman@21776: norm h * KF * norm (g (x + h) - g x) * K) / norm h" huffman@21776: by (intro huffman@21776: divide_right_mono mult_mono' huffman@21776: order_trans [OF norm_triangle_ineq add_mono] huffman@21776: order_trans [OF norm_prod mult_right_mono] huffman@21776: mult_nonneg_nonneg order_refl norm_ge_zero norm_F huffman@21776: K [THEN order_less_imp_le] huffman@21776: ) huffman@21776: also have "\ = ?fun2 h" huffman@21776: by (simp add: add_divide_distrib) huffman@21776: finally show "?fun1 h \ ?fun2 h" . huffman@21776: qed huffman@21776: thus "(\h. huffman@21776: norm (f (x + h) ** g (x + h) - f x ** g x - (f x ** G h + F h ** g x)) huffman@21776: / norm h) -- 0 --> 0" huffman@21776: by (simp only: FDERIV_lemma) huffman@21776: qed huffman@21776: huffman@44282: lemmas FDERIV_mult = huffman@44282: bounded_bilinear.FDERIV [OF bounded_bilinear_mult] huffman@21776: huffman@44282: lemmas FDERIV_scaleR = huffman@44282: bounded_bilinear.FDERIV [OF bounded_bilinear_scaleR] wenzelm@28866: huffman@21776: huffman@21776: subsection {* Powers *} huffman@21776: huffman@21776: lemma FDERIV_power_Suc: haftmann@31021: fixes x :: "'a::{real_normed_algebra,comm_ring_1}" wenzelm@28866: shows "FDERIV (\x. x ^ Suc n) x :> (\h. (1 + of_nat n) * x ^ n * h)" huffman@21776: apply (induct n) huffman@36361: apply (simp add: FDERIV_ident) huffman@21776: apply (drule FDERIV_mult [OF FDERIV_ident]) wenzelm@28866: apply (simp only: of_nat_Suc left_distrib mult_1_left) wenzelm@28866: apply (simp only: power_Suc right_distrib add_ac mult_ac) huffman@21776: done huffman@21776: huffman@21776: lemma FDERIV_power: haftmann@31021: fixes x :: "'a::{real_normed_algebra,comm_ring_1}" huffman@21776: shows "FDERIV (\x. x ^ n) x :> (\h. of_nat n * x ^ (n - 1) * h)" wenzelm@28866: apply (cases n) wenzelm@28866: apply (simp add: FDERIV_const) huffman@30273: apply (simp add: FDERIV_power_Suc del: power_Suc) wenzelm@28866: done wenzelm@28866: huffman@21776: huffman@21776: subsection {* Inverse *} huffman@21776: huffman@21776: lemmas bounded_linear_mult_const = huffman@44282: bounded_linear_mult_left [THEN bounded_linear_compose] huffman@21776: huffman@21776: lemmas bounded_linear_const_mult = huffman@44282: bounded_linear_mult_right [THEN bounded_linear_compose] huffman@21776: huffman@21776: lemma FDERIV_inverse: huffman@21776: fixes x :: "'a::real_normed_div_algebra" huffman@21776: assumes x: "x \ 0" huffman@21776: shows "FDERIV inverse x :> (\h. - (inverse x * h * inverse x))" huffman@21776: (is "FDERIV ?inv _ :> _") huffman@21776: proof (rule FDERIV_I) huffman@21776: show "bounded_linear (\h. - (?inv x * h * ?inv x))" huffman@21776: apply (rule bounded_linear_minus) huffman@21776: apply (rule bounded_linear_mult_const) huffman@21776: apply (rule bounded_linear_const_mult) huffman@21776: apply (rule bounded_linear_ident) huffman@21776: done huffman@21776: next huffman@21776: show "(\h. norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h) huffman@21776: -- 0 --> 0" huffman@21776: proof (rule LIM_equal2) huffman@21776: show "0 < norm x" using x by simp huffman@21776: next huffman@21776: fix h::'a huffman@21776: assume 1: "h \ 0" huffman@21776: assume "norm (h - 0) < norm x" huffman@21776: hence "h \ -x" by clarsimp huffman@21776: hence 2: "x + h \ 0" huffman@21776: apply (rule contrapos_nn) huffman@21776: apply (rule sym) huffman@34146: apply (erule minus_unique) huffman@21776: done huffman@21776: show "norm (?inv (x + h) - ?inv x - - (?inv x * h * ?inv x)) / norm h huffman@21776: = norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" huffman@21776: apply (subst inverse_diff_inverse [OF 2 x]) huffman@21776: apply (subst minus_diff_minus) huffman@21776: apply (subst norm_minus_cancel) huffman@21776: apply (simp add: left_diff_distrib) huffman@21776: done huffman@21776: next huffman@21776: show "(\h. norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h) huffman@21776: -- 0 --> 0" huffman@21776: proof (rule real_LIM_sandwich_zero) huffman@21776: show "(\h. norm (?inv (x + h) - ?inv x) * norm (?inv x)) huffman@21776: -- 0 --> 0" huffman@44568: apply (rule tendsto_mult_left_zero) huffman@44568: apply (rule tendsto_norm_zero) huffman@21776: apply (rule LIM_zero) huffman@21776: apply (rule LIM_offset_zero) huffman@44568: apply (rule tendsto_inverse) huffman@44568: apply (rule tendsto_ident_at) huffman@21776: apply (rule x) huffman@21776: done huffman@21776: next huffman@21776: fix h::'a assume h: "h \ 0" huffman@21776: show "0 \ norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h" huffman@21776: apply (rule divide_nonneg_pos) huffman@21776: apply (rule norm_ge_zero) huffman@21776: apply (simp add: h) huffman@21776: done huffman@21776: next huffman@21776: fix h::'a assume h: "h \ 0" huffman@21776: have "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h huffman@21776: \ norm (?inv (x + h) - ?inv x) * norm h * norm (?inv x) / norm h" huffman@21776: apply (rule divide_right_mono [OF _ norm_ge_zero]) huffman@21776: apply (rule order_trans [OF norm_mult_ineq]) huffman@21776: apply (rule mult_right_mono [OF _ norm_ge_zero]) huffman@21776: apply (rule norm_mult_ineq) huffman@21776: done huffman@21776: also have "\ = norm (?inv (x + h) - ?inv x) * norm (?inv x)" huffman@21776: by simp huffman@21776: finally show "norm ((?inv (x + h) - ?inv x) * h * ?inv x) / norm h huffman@37729: \ norm (?inv (x + h) - ?inv x) * norm (?inv x)" . huffman@21776: qed huffman@21776: qed huffman@21776: qed huffman@21776: huffman@21776: subsection {* Alternate definition *} huffman@21776: huffman@21776: lemma field_fderiv_def: huffman@21776: fixes x :: "'a::real_normed_field" shows huffman@21776: "FDERIV f x :> (\h. h * D) = (\h. (f (x + h) - f x) / h) -- 0 --> D" huffman@21776: apply (unfold fderiv_def) huffman@44282: apply (simp add: bounded_linear_mult_left) huffman@21776: apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) huffman@21776: apply (subst diff_divide_distrib) huffman@21776: apply (subst times_divide_eq_left [symmetric]) nipkow@23398: apply (simp cong: LIM_cong) huffman@44568: apply (simp add: tendsto_norm_zero_iff LIM_zero_iff) huffman@21776: done huffman@21776: huffman@21776: end