huffman@21164: (* Title : Deriv.thy huffman@21164: Author : Jacques D. Fleuriot huffman@21164: Copyright : 1998 University of Cambridge hoelzl@51642: Author : Brian Huffman huffman@21164: Conversion to Isar and new proofs by Lawrence C Paulson, 2004 huffman@21164: GMVT by Benjamin Porter, 2005 huffman@21164: *) huffman@21164: huffman@21164: header{* Differentiation *} huffman@21164: huffman@21164: theory Deriv hoelzl@51526: imports Limits huffman@21164: begin huffman@21164: hoelzl@56181: subsection {* Frechet derivative *} hoelzl@56181: hoelzl@51642: definition hoelzl@51642: has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" hoelzl@56182: (infix "(has'_derivative)" 50) hoelzl@51642: where hoelzl@51642: "(f has_derivative f') F \ hoelzl@51642: (bounded_linear f' \ hoelzl@51642: ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) ---> 0) F)" hoelzl@51642: hoelzl@56181: text {* hoelzl@56181: Usually the filter @{term F} is @{term "at x within s"}. @{term "(f has_derivative D) hoelzl@56181: (at x within s)"} means: @{term D} is the derivative of function @{term f} at point @{term x} hoelzl@56181: within the set @{term s}. Where @{term s} is used to express left or right sided derivatives. In hoelzl@56181: most cases @{term s} is either a variable or @{term UNIV}. hoelzl@56181: *} hoelzl@56181: hoelzl@56381: lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" hoelzl@56381: by simp hoelzl@56381: hoelzl@56381: definition hoelzl@56381: has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" hoelzl@56381: (infix "(has'_field'_derivative)" 50) hoelzl@56381: where hoelzl@56381: "(f has_field_derivative D) F \ (f has_derivative op * D) F" hoelzl@56381: hoelzl@56381: lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" hoelzl@56381: by simp hoelzl@56381: hoelzl@56381: definition hoelzl@56381: has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" hoelzl@56381: (infix "has'_vector'_derivative" 50) hoelzl@56381: where hoelzl@56381: "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" hoelzl@56381: hoelzl@56381: lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" hoelzl@56381: by simp hoelzl@56381: hoelzl@56381: ML {* hoelzl@56381: hoelzl@56381: structure Derivative_Intros = Named_Thms hoelzl@56381: ( hoelzl@56381: val name = @{binding derivative_intros} hoelzl@56381: val description = "structural introduction rules for derivatives" hoelzl@56381: ) hoelzl@56381: hoelzl@56381: *} hoelzl@56381: hoelzl@56381: setup {* hoelzl@56381: let hoelzl@56381: val eq_thms = [@{thm has_derivative_eq_rhs}, @{thm DERIV_cong}, @{thm has_vector_derivative_eq_rhs}] hoelzl@56381: fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms hoelzl@56381: in hoelzl@56381: Derivative_Intros.setup #> hoelzl@56381: Global_Theory.add_thms_dynamic hoelzl@56381: (@{binding derivative_eq_intros}, map_filter eq_rule o Derivative_Intros.get o Context.proof_of) hoelzl@56381: end; hoelzl@56381: *} hoelzl@56381: hoelzl@56181: text {* hoelzl@56181: The following syntax is only used as a legacy syntax. hoelzl@56181: *} hoelzl@56181: abbreviation (input) hoelzl@56181: FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" hoelzl@56181: ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) hoelzl@56181: where hoelzl@56181: "FDERIV f x :> f' \ (f has_derivative f') (at x)" hoelzl@56181: hoelzl@56181: lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" hoelzl@51642: by (simp add: has_derivative_def) hoelzl@51642: hoelzl@56369: lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" hoelzl@56369: using bounded_linear.linear[OF has_derivative_bounded_linear] . hoelzl@56369: hoelzl@56381: lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" hoelzl@51642: by (simp add: has_derivative_def tendsto_const) hoelzl@51642: hoelzl@56381: lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" hoelzl@51642: by (simp add: has_derivative_def tendsto_const) hoelzl@51642: hoelzl@51642: lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. hoelzl@51642: hoelzl@56181: lemma (in bounded_linear) has_derivative: hoelzl@51642: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" hoelzl@51642: using assms unfolding has_derivative_def hoelzl@51642: apply safe wenzelm@56219: apply (erule bounded_linear_compose [OF bounded_linear]) wenzelm@56219: apply (drule tendsto) wenzelm@56219: apply (simp add: scaleR diff add zero) hoelzl@51642: done hoelzl@51642: hoelzl@56381: lemmas has_derivative_scaleR_right [derivative_intros] = hoelzl@56181: bounded_linear.has_derivative [OF bounded_linear_scaleR_right] hoelzl@51642: hoelzl@56381: lemmas has_derivative_scaleR_left [derivative_intros] = hoelzl@56181: bounded_linear.has_derivative [OF bounded_linear_scaleR_left] hoelzl@51642: hoelzl@56381: lemmas has_derivative_mult_right [derivative_intros] = hoelzl@56181: bounded_linear.has_derivative [OF bounded_linear_mult_right] hoelzl@51642: hoelzl@56381: lemmas has_derivative_mult_left [derivative_intros] = hoelzl@56181: bounded_linear.has_derivative [OF bounded_linear_mult_left] hoelzl@51642: hoelzl@56381: lemma has_derivative_add[simp, derivative_intros]: hoelzl@51642: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" hoelzl@51642: shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" hoelzl@51642: unfolding has_derivative_def hoelzl@51642: proof safe hoelzl@51642: let ?x = "Lim F (\x. x)" hoelzl@51642: let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" hoelzl@51642: have "((\x. ?D f f' x + ?D g g' x) ---> (0 + 0)) F" hoelzl@51642: using f g by (intro tendsto_add) (auto simp: has_derivative_def) hoelzl@51642: then show "(?D (\x. f x + g x) (\x. f' x + g' x) ---> 0) F" hoelzl@51642: by (simp add: field_simps scaleR_add_right scaleR_diff_right) hoelzl@56181: qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) hoelzl@51642: hoelzl@56381: lemma has_derivative_setsum[simp, derivative_intros]: hoelzl@51642: assumes f: "\i. i \ I \ (f i has_derivative f' i) F" hoelzl@51642: shows "((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" hoelzl@51642: proof cases hoelzl@51642: assume "finite I" from this f show ?thesis hoelzl@51642: by induct (simp_all add: f) hoelzl@51642: qed simp hoelzl@51642: hoelzl@56381: lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" hoelzl@56181: using has_derivative_scaleR_right[of f f' F "-1"] by simp hoelzl@51642: hoelzl@56381: lemma has_derivative_diff[simp, derivative_intros]: hoelzl@56181: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" hoelzl@56181: by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) hoelzl@51642: hoelzl@56181: lemma has_derivative_at_within: hoelzl@56181: "(f has_derivative f') (at x within s) \ hoelzl@51642: (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) ---> 0) (at x within s))" hoelzl@51642: by (cases "at x within s = bot") (simp_all add: has_derivative_def Lim_ident_at) hoelzl@51642: hoelzl@56181: lemma has_derivative_iff_norm: hoelzl@56181: "(f has_derivative f') (at x within s) \ hoelzl@51642: (bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) ---> 0) (at x within s))" hoelzl@51642: using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] hoelzl@56181: by (simp add: has_derivative_at_within divide_inverse ac_simps) hoelzl@51642: hoelzl@56181: lemma has_derivative_at: hoelzl@56181: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" hoelzl@56181: unfolding has_derivative_iff_norm LIM_offset_zero_iff[of _ _ x] by simp hoelzl@51642: hoelzl@56181: lemma field_has_derivative_at: hoelzl@51642: fixes x :: "'a::real_normed_field" hoelzl@56181: shows "(f has_derivative op * D) (at x) \ (\h. (f (x + h) - f x) / h) -- 0 --> D" hoelzl@56181: apply (unfold has_derivative_at) hoelzl@56181: apply (simp add: bounded_linear_mult_right) hoelzl@51642: apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) hoelzl@51642: apply (subst diff_divide_distrib) hoelzl@51642: apply (subst times_divide_eq_left [symmetric]) hoelzl@51642: apply (simp cong: LIM_cong) hoelzl@51642: apply (simp add: tendsto_norm_zero_iff LIM_zero_iff) hoelzl@51642: done hoelzl@51642: hoelzl@56181: lemma has_derivativeI: hoelzl@51642: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) ---> 0) (at x within s) \ hoelzl@56181: (f has_derivative f') (at x within s)" hoelzl@56181: by (simp add: has_derivative_at_within) hoelzl@51642: hoelzl@56181: lemma has_derivativeI_sandwich: hoelzl@51642: assumes e: "0 < e" and bounded: "bounded_linear f'" hoelzl@51642: and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" hoelzl@51642: and "(H ---> 0) (at x within s)" hoelzl@56181: shows "(f has_derivative f') (at x within s)" hoelzl@56181: unfolding has_derivative_iff_norm hoelzl@51642: proof safe hoelzl@51642: show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) ---> 0) (at x within s)" hoelzl@51642: proof (rule tendsto_sandwich[where f="\x. 0"]) hoelzl@51642: show "(H ---> 0) (at x within s)" by fact hoelzl@51642: show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" hoelzl@51642: unfolding eventually_at using e sandwich by auto hoelzl@51642: qed (auto simp: le_divide_eq tendsto_const) hoelzl@51642: qed fact hoelzl@51642: hoelzl@56181: lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" hoelzl@56181: by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) hoelzl@51642: lp15@56261: lemmas has_derivative_within_subset = has_derivative_subset lp15@56261: lp15@56261: hoelzl@51642: subsection {* Continuity *} hoelzl@51642: hoelzl@56181: lemma has_derivative_continuous: hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" hoelzl@51642: shows "continuous (at x within s) f" hoelzl@51642: proof - hoelzl@56181: from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) hoelzl@51642: note F.tendsto[tendsto_intros] hoelzl@51642: let ?L = "\f. (f ---> 0) (at x within s)" hoelzl@51642: have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" hoelzl@56181: using f unfolding has_derivative_iff_norm by blast hoelzl@51642: then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) hoelzl@51642: by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) hoelzl@51642: also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" hoelzl@51642: by (intro filterlim_cong) (simp_all add: eventually_at_filter) hoelzl@51642: finally have "?L (\y. (f y - f x) - f' (y - x))" hoelzl@51642: by (rule tendsto_norm_zero_cancel) hoelzl@51642: then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" hoelzl@51642: by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) hoelzl@51642: then have "?L (\y. f y - f x)" hoelzl@51642: by simp hoelzl@51642: from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis hoelzl@51642: by (simp add: continuous_within) hoelzl@51642: qed hoelzl@51642: hoelzl@51642: subsection {* Composition *} hoelzl@51642: hoelzl@51642: lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f ---> y) (at x within s) \ (f ---> y) (inf (nhds x) (principal s))" hoelzl@51642: unfolding tendsto_def eventually_inf_principal eventually_at_filter hoelzl@51642: by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) hoelzl@51642: hoelzl@56181: lemma has_derivative_in_compose: hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" hoelzl@56181: assumes g: "(g has_derivative g') (at (f x) within (f`s))" hoelzl@56181: shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" hoelzl@51642: proof - hoelzl@56181: from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) hoelzl@56181: from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) hoelzl@51642: from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast hoelzl@51642: from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast hoelzl@51642: note G.tendsto[tendsto_intros] hoelzl@51642: hoelzl@51642: let ?L = "\f. (f ---> 0) (at x within s)" hoelzl@51642: let ?D = "\f f' x y. (f y - f x) - f' (y - x)" hoelzl@51642: let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" hoelzl@51642: let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" hoelzl@51642: def Nf \ "?N f f' x" hoelzl@51642: def Ng \ "\y. ?N g g' (f x) (f y)" hoelzl@51642: hoelzl@51642: show ?thesis hoelzl@56181: proof (rule has_derivativeI_sandwich[of 1]) hoelzl@51642: show "bounded_linear (\x. g' (f' x))" hoelzl@56181: using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) hoelzl@51642: next hoelzl@51642: fix y::'a assume neq: "y \ x" hoelzl@51642: have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" hoelzl@51642: by (simp add: G.diff G.add field_simps) hoelzl@51642: also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" hoelzl@51642: by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) hoelzl@51642: also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" hoelzl@51642: proof (intro add_mono mult_left_mono) hoelzl@51642: have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" hoelzl@51642: by simp hoelzl@51642: also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" hoelzl@51642: by (rule norm_triangle_ineq) hoelzl@51642: also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" hoelzl@51642: using kF by (intro add_mono) simp hoelzl@51642: finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" hoelzl@51642: by (simp add: neq Nf_def field_simps) hoelzl@51642: qed (insert kG, simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps) hoelzl@51642: finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . hoelzl@51642: next hoelzl@51642: have [tendsto_intros]: "?L Nf" hoelzl@56181: using f unfolding has_derivative_iff_norm Nf_def .. hoelzl@51642: from f have "(f ---> f x) (at x within s)" hoelzl@56181: by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) hoelzl@51642: then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" hoelzl@51642: unfolding filterlim_def hoelzl@51642: by (simp add: eventually_filtermap eventually_at_filter le_principal) hoelzl@51642: hoelzl@51642: have "((?N g g' (f x)) ---> 0) (at (f x) within f`s)" hoelzl@56181: using g unfolding has_derivative_iff_norm .. hoelzl@51642: then have g': "((?N g g' (f x)) ---> 0) (inf (nhds (f x)) (principal (f`s)))" hoelzl@51642: by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp hoelzl@51642: hoelzl@51642: have [tendsto_intros]: "?L Ng" hoelzl@51642: unfolding Ng_def by (rule filterlim_compose[OF g' f']) hoelzl@51642: show "((\y. Nf y * kG + Ng y * (Nf y + kF)) ---> 0) (at x within s)" hoelzl@51642: by (intro tendsto_eq_intros) auto hoelzl@51642: qed simp hoelzl@51642: qed hoelzl@51642: hoelzl@56181: lemma has_derivative_compose: hoelzl@56181: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ hoelzl@56181: ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" hoelzl@56181: by (blast intro: has_derivative_in_compose has_derivative_subset) hoelzl@51642: hoelzl@51642: lemma (in bounded_bilinear) FDERIV: hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" hoelzl@56181: shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" hoelzl@51642: proof - hoelzl@56181: from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] hoelzl@51642: obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast hoelzl@51642: hoelzl@51642: from pos_bounded obtain K where K: "0 < K" and norm_prod: hoelzl@51642: "\a b. norm (a ** b) \ norm a * norm b * K" by fast hoelzl@51642: let ?D = "\f f' y. f y - f x - f' (y - x)" hoelzl@51642: let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" hoelzl@51642: def Ng =="?N g g'" and Nf =="?N f f'" hoelzl@51642: hoelzl@51642: let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" hoelzl@51642: let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" hoelzl@51642: let ?F = "at x within s" huffman@21164: hoelzl@51642: show ?thesis hoelzl@56181: proof (rule has_derivativeI_sandwich[of 1]) hoelzl@51642: show "bounded_linear (\h. f x ** g' h + f' h ** g x)" hoelzl@51642: by (intro bounded_linear_add hoelzl@51642: bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] hoelzl@56181: has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) hoelzl@51642: next hoelzl@51642: from g have "(g ---> g x) ?F" hoelzl@56181: by (intro continuous_within[THEN iffD1] has_derivative_continuous) hoelzl@51642: moreover from f g have "(Nf ---> 0) ?F" "(Ng ---> 0) ?F" hoelzl@56181: by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) hoelzl@51642: ultimately have "(?fun2 ---> norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" hoelzl@51642: by (intro tendsto_intros) (simp_all add: LIM_zero_iff) hoelzl@51642: then show "(?fun2 ---> 0) ?F" hoelzl@51642: by simp hoelzl@51642: next hoelzl@51642: fix y::'d assume "y \ x" hoelzl@51642: have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" hoelzl@51642: by (simp add: diff_left diff_right add_left add_right field_simps) hoelzl@51642: also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + hoelzl@51642: norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" hoelzl@51642: by (intro divide_right_mono mult_mono' hoelzl@51642: order_trans [OF norm_triangle_ineq add_mono] hoelzl@51642: order_trans [OF norm_prod mult_right_mono] hoelzl@51642: mult_nonneg_nonneg order_refl norm_ge_zero norm_F hoelzl@51642: K [THEN order_less_imp_le]) hoelzl@51642: also have "\ = ?fun2 y" hoelzl@51642: by (simp add: add_divide_distrib Ng_def Nf_def) hoelzl@51642: finally show "?fun1 y \ ?fun2 y" . hoelzl@51642: qed simp hoelzl@51642: qed hoelzl@51642: hoelzl@56381: lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] hoelzl@56381: lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] hoelzl@51642: hoelzl@56381: lemma has_derivative_setprod[simp, derivative_intros]: hoelzl@51642: fixes f :: "'i \ 'a :: real_normed_vector \ 'b :: real_normed_field" hoelzl@56181: assumes f: "\i. i \ I \ (f i has_derivative f' i) (at x within s)" hoelzl@56181: shows "((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within s)" hoelzl@51642: proof cases hoelzl@51642: assume "finite I" from this f show ?thesis hoelzl@51642: proof induct hoelzl@51642: case (insert i I) hoelzl@51642: let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" hoelzl@56181: have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within s)" hoelzl@56181: using insert by (intro has_derivative_mult) auto hoelzl@51642: also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" haftmann@57418: using insert(1,2) by (auto simp add: setsum_right_distrib insert_Diff_if intro!: ext setsum.cong) hoelzl@51642: finally show ?case hoelzl@51642: using insert by simp hoelzl@51642: qed simp hoelzl@51642: qed simp hoelzl@51642: hoelzl@56381: lemma has_derivative_power[simp, derivative_intros]: hoelzl@51642: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" hoelzl@56181: shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within s)" hoelzl@56181: using has_derivative_setprod[OF f, of "{..< n}"] by (simp add: setprod_constant ac_simps) hoelzl@51642: hoelzl@56181: lemma has_derivative_inverse': hoelzl@51642: fixes x :: "'a::real_normed_div_algebra" hoelzl@51642: assumes x: "x \ 0" hoelzl@56181: shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within s)" hoelzl@56181: (is "(?inv has_derivative ?f) _") hoelzl@56181: proof (rule has_derivativeI_sandwich) hoelzl@51642: show "bounded_linear (\h. - (?inv x * h * ?inv x))" hoelzl@51642: apply (rule bounded_linear_minus) hoelzl@51642: apply (rule bounded_linear_mult_const) hoelzl@51642: apply (rule bounded_linear_const_mult) hoelzl@51642: apply (rule bounded_linear_ident) hoelzl@51642: done hoelzl@51642: next hoelzl@51642: show "0 < norm x" using x by simp hoelzl@51642: next hoelzl@51642: show "((\y. norm (?inv y - ?inv x) * norm (?inv x)) ---> 0) (at x within s)" hoelzl@51642: apply (rule tendsto_mult_left_zero) hoelzl@51642: apply (rule tendsto_norm_zero) hoelzl@51642: apply (rule LIM_zero) hoelzl@51642: apply (rule tendsto_inverse) hoelzl@51642: apply (rule tendsto_ident_at) hoelzl@51642: apply (rule x) hoelzl@51642: done hoelzl@51642: next hoelzl@51642: fix y::'a assume h: "y \ x" "dist y x < norm x" hoelzl@51642: then have "y \ 0" hoelzl@51642: by (auto simp: norm_conv_dist dist_commute) hoelzl@51642: have "norm (?inv y - ?inv x - ?f (y -x)) / norm (y - x) = norm ((?inv y - ?inv x) * (y - x) * ?inv x) / norm (y - x)" hoelzl@51642: apply (subst inverse_diff_inverse [OF `y \ 0` x]) hoelzl@51642: apply (subst minus_diff_minus) hoelzl@51642: apply (subst norm_minus_cancel) hoelzl@51642: apply (simp add: left_diff_distrib) hoelzl@51642: done hoelzl@51642: also have "\ \ norm (?inv y - ?inv x) * norm (y - x) * norm (?inv x) / norm (y - x)" hoelzl@51642: apply (rule divide_right_mono [OF _ norm_ge_zero]) hoelzl@51642: apply (rule order_trans [OF norm_mult_ineq]) hoelzl@51642: apply (rule mult_right_mono [OF _ norm_ge_zero]) hoelzl@51642: apply (rule norm_mult_ineq) hoelzl@51642: done hoelzl@51642: also have "\ = norm (?inv y - ?inv x) * norm (?inv x)" hoelzl@51642: by simp hoelzl@51642: finally show "norm (?inv y - ?inv x - ?f (y -x)) / norm (y - x) \ hoelzl@51642: norm (?inv y - ?inv x) * norm (?inv x)" . hoelzl@51642: qed hoelzl@51642: hoelzl@56381: lemma has_derivative_inverse[simp, derivative_intros]: hoelzl@51642: fixes f :: "_ \ 'a::real_normed_div_algebra" hoelzl@56181: assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within s)" hoelzl@56181: shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within s)" hoelzl@56181: using has_derivative_compose[OF f has_derivative_inverse', OF x] . hoelzl@51642: hoelzl@56381: lemma has_derivative_divide[simp, derivative_intros]: hoelzl@51642: fixes f :: "_ \ 'a::real_normed_div_algebra" hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" lp15@55967: assumes x: "g x \ 0" hoelzl@56181: shows "((\x. f x / g x) has_derivative hoelzl@56181: (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within s)" hoelzl@56181: using has_derivative_mult[OF f has_derivative_inverse[OF x g]] hoelzl@56480: by (simp add: field_simps) lp15@55967: lp15@55967: text{*Conventional form requires mult-AC laws. Types real and complex only.*} hoelzl@56181: hoelzl@56381: lemma has_derivative_divide'[derivative_intros]: lp15@55967: fixes f :: "_ \ 'a::real_normed_field" hoelzl@56181: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" and x: "g x \ 0" hoelzl@56181: shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within s)" lp15@55967: proof - lp15@55967: { fix h lp15@55967: have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = lp15@55967: (f' h * g x - f x * g' h) / (g x * g x)" hoelzl@56480: by (simp add: field_simps x) lp15@55967: } lp15@55967: then show ?thesis hoelzl@56181: using has_derivative_divide [OF f g] x lp15@55967: by simp lp15@55967: qed hoelzl@51642: hoelzl@51642: subsection {* Uniqueness *} hoelzl@51642: hoelzl@51642: text {* hoelzl@51642: hoelzl@56181: This can not generally shown for @{const has_derivative}, as we need to approach the point from hoelzl@51642: all directions. There is a proof in @{text Multivariate_Analysis} for @{text euclidean_space}. hoelzl@51642: hoelzl@51642: *} hoelzl@51642: hoelzl@56181: lemma has_derivative_zero_unique: hoelzl@56181: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" hoelzl@51642: proof - hoelzl@51642: interpret F: bounded_linear F hoelzl@56181: using assms by (rule has_derivative_bounded_linear) hoelzl@51642: let ?r = "\h. norm (F h) / norm h" hoelzl@51642: have *: "?r -- 0 --> 0" hoelzl@56181: using assms unfolding has_derivative_at by simp hoelzl@51642: show "F = (\h. 0)" hoelzl@51642: proof hoelzl@51642: fix h show "F h = 0" hoelzl@51642: proof (rule ccontr) wenzelm@53374: assume **: "F h \ 0" nipkow@56541: hence h: "h \ 0" by (clarsimp simp add: F.zero) nipkow@56541: with ** have "0 < ?r h" by simp hoelzl@51642: from LIM_D [OF * this] obtain s where s: "0 < s" hoelzl@51642: and r: "\x. x \ 0 \ norm x < s \ ?r x < ?r h" by auto hoelzl@51642: from dense [OF s] obtain t where t: "0 < t \ t < s" .. hoelzl@51642: let ?x = "scaleR (t / norm h) h" hoelzl@51642: have "?x \ 0" and "norm ?x < s" using t h by simp_all hoelzl@51642: hence "?r ?x < ?r h" by (rule r) hoelzl@51642: thus "False" using t h by (simp add: F.scaleR) hoelzl@51642: qed hoelzl@51642: qed hoelzl@51642: qed hoelzl@51642: hoelzl@56181: lemma has_derivative_unique: hoelzl@56181: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" hoelzl@51642: proof - hoelzl@56181: have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" hoelzl@56181: using has_derivative_diff [OF assms] by simp hoelzl@51642: hence "(\h. F h - F' h) = (\h. 0)" hoelzl@56181: by (rule has_derivative_zero_unique) hoelzl@51642: thus "F = F'" hoelzl@51642: unfolding fun_eq_iff right_minus_eq . hoelzl@51642: qed hoelzl@51642: hoelzl@51642: subsection {* Differentiability predicate *} hoelzl@51642: hoelzl@56181: definition hoelzl@56181: differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" hoelzl@56182: (infix "differentiable" 50) hoelzl@56181: where hoelzl@56181: "f differentiable F \ (\D. (f has_derivative D) F)" hoelzl@51642: hoelzl@56181: lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_subset) hoelzl@51642: lp15@56261: lemmas differentiable_within_subset = differentiable_subset lp15@56261: hoelzl@56381: lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_ident) hoelzl@51642: hoelzl@56381: lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_const) hoelzl@51642: hoelzl@51642: lemma differentiable_in_compose: hoelzl@56181: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_in_compose) hoelzl@51642: hoelzl@51642: lemma differentiable_compose: hoelzl@56181: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" hoelzl@51642: by (blast intro: differentiable_in_compose differentiable_subset) hoelzl@51642: hoelzl@56381: lemma differentiable_sum [simp, derivative_intros]: hoelzl@56181: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_add) hoelzl@51642: hoelzl@56381: lemma differentiable_minus [simp, derivative_intros]: hoelzl@56181: "f differentiable F \ (\x. - f x) differentiable F" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_minus) hoelzl@51642: hoelzl@56381: lemma differentiable_diff [simp, derivative_intros]: hoelzl@56181: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_diff) hoelzl@51642: hoelzl@56381: lemma differentiable_mult [simp, derivative_intros]: hoelzl@51642: fixes f g :: "'a :: real_normed_vector \ 'b :: real_normed_algebra" hoelzl@56181: shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_mult) hoelzl@51642: hoelzl@56381: lemma differentiable_inverse [simp, derivative_intros]: hoelzl@51642: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" hoelzl@56181: shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_inverse) hoelzl@51642: hoelzl@56381: lemma differentiable_divide [simp, derivative_intros]: hoelzl@51642: fixes f g :: "'a :: real_normed_vector \ 'b :: real_normed_field" hoelzl@56181: shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" hoelzl@51642: unfolding divide_inverse using assms by simp hoelzl@51642: hoelzl@56381: lemma differentiable_power [simp, derivative_intros]: hoelzl@51642: fixes f g :: "'a :: real_normed_vector \ 'b :: real_normed_field" hoelzl@56181: shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_power) hoelzl@51642: hoelzl@56381: lemma differentiable_scaleR [simp, derivative_intros]: hoelzl@56181: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" hoelzl@56181: unfolding differentiable_def by (blast intro: has_derivative_scaleR) hoelzl@51642: hoelzl@56181: lemma has_derivative_imp_has_field_derivative: hoelzl@56181: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" hoelzl@56181: unfolding has_field_derivative_def haftmann@57512: by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) hoelzl@56181: hoelzl@56181: lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative op * D) F" hoelzl@56181: by (simp add: has_field_derivative_def) hoelzl@51642: lp15@56261: lemma DERIV_subset: lp15@56261: "(f has_field_derivative f') (at x within s) \ t \ s lp15@56261: \ (f has_field_derivative f') (at x within t)" lp15@56261: by (simp add: has_field_derivative_def has_derivative_within_subset) lp15@56261: hoelzl@56181: abbreviation (input) hoelzl@56381: DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" hoelzl@56181: ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) hoelzl@51642: where hoelzl@56181: "DERIV f x :> D \ (f has_field_derivative D) (at x)" hoelzl@51642: hoelzl@56181: abbreviation hoelzl@56181: has_real_derivative :: "(real \ real) \ real \ real filter \ bool" hoelzl@56182: (infix "(has'_real'_derivative)" 50) hoelzl@56181: where hoelzl@56181: "(f has_real_derivative D) F \ (f has_field_derivative D) F" hoelzl@56181: hoelzl@56181: lemma real_differentiable_def: hoelzl@56181: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" hoelzl@51642: proof safe hoelzl@56181: assume "f differentiable at x within s" hoelzl@56181: then obtain f' where *: "(f has_derivative f') (at x within s)" hoelzl@56181: unfolding differentiable_def by auto hoelzl@56181: then obtain c where "f' = (op * c)" haftmann@57512: by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) hoelzl@56181: with * show "\D. (f has_real_derivative D) (at x within s)" hoelzl@56181: unfolding has_field_derivative_def by auto hoelzl@56181: qed (auto simp: differentiable_def has_field_derivative_def) hoelzl@51642: hoelzl@56181: lemma real_differentiableE [elim?]: hoelzl@56181: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" hoelzl@56181: using assms by (auto simp: real_differentiable_def) hoelzl@51642: hoelzl@56181: lemma differentiableD: "f differentiable (at x within s) \ \D. (f has_real_derivative D) (at x within s)" hoelzl@56181: by (auto elim: real_differentiableE) hoelzl@51642: hoelzl@56181: lemma differentiableI: "(f has_real_derivative D) (at x within s) \ f differentiable (at x within s)" hoelzl@56181: by (force simp add: real_differentiable_def) hoelzl@51642: hoelzl@56381: lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) -- 0 --> D" hoelzl@56181: apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right LIM_zero_iff[symmetric, of _ D]) hoelzl@51642: apply (subst (2) tendsto_norm_zero_iff[symmetric]) hoelzl@51642: apply (rule filterlim_cong) hoelzl@51642: apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) hoelzl@51642: done huffman@21164: hoelzl@56181: lemma mult_commute_abs: "(\x. x * c) = op * (c::'a::ab_semigroup_mult)" haftmann@57512: by (simp add: fun_eq_iff mult.commute) huffman@21164: hoelzl@56181: subsection {* Derivatives *} huffman@21164: hoelzl@51642: lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) -- 0 --> D" hoelzl@56381: by (simp add: DERIV_def) huffman@21164: hoelzl@56381: lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto hoelzl@56181: hoelzl@56381: lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto huffman@21164: hoelzl@56381: lemma field_differentiable_add[derivative_intros]: hoelzl@56381: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ hoelzl@56381: ((\z. f z + g z) has_field_derivative f' + g') F" hoelzl@56381: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) hoelzl@56381: (auto simp: has_field_derivative_def field_simps mult_commute_abs) lp15@56261: lp15@56261: corollary DERIV_add: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ hoelzl@56181: ((\x. f x + g x) has_field_derivative D + E) (at x within s)" lp15@56261: by (rule field_differentiable_add) lp15@56261: hoelzl@56381: lemma field_differentiable_minus[derivative_intros]: hoelzl@56381: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" hoelzl@56381: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) hoelzl@56381: (auto simp: has_field_derivative_def field_simps mult_commute_abs) huffman@21164: lp15@56261: corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" lp15@56261: by (rule field_differentiable_minus) huffman@21164: hoelzl@56381: lemma field_differentiable_diff[derivative_intros]: hoelzl@56381: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" hoelzl@56381: by (simp only: assms diff_conv_add_uminus field_differentiable_add field_differentiable_minus) lp15@56261: lp15@56261: corollary DERIV_diff: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ hoelzl@56181: ((\x. f x - g x) has_field_derivative D - E) (at x within s)" lp15@56261: by (rule field_differentiable_diff) hoelzl@51642: hoelzl@56181: lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" hoelzl@56181: by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp huffman@21164: lp15@56261: corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" lp15@56261: by (rule DERIV_continuous) lp15@56261: lp15@56261: lemma DERIV_continuous_on: lp15@56261: "(\x. x \ s \ (f has_field_derivative D) (at x)) \ continuous_on s f" lp15@56261: by (metis DERIV_continuous continuous_at_imp_continuous_on) hoelzl@51642: hoelzl@56181: lemma DERIV_mult': hoelzl@56181: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ hoelzl@56181: ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) hoelzl@56181: (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) huffman@21164: hoelzl@56381: lemma DERIV_mult[derivative_intros]: hoelzl@56181: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ hoelzl@56181: ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) hoelzl@56181: (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) hoelzl@51642: hoelzl@51642: text {* Derivative of linear multiplication *} huffman@21164: hoelzl@51642: lemma DERIV_cmult: hoelzl@56181: "(f has_field_derivative D) (at x within s) ==> ((\x. c * f x) has_field_derivative c * D) (at x within s)" hoelzl@51642: by (drule DERIV_mult' [OF DERIV_const], simp) huffman@21164: lp15@55967: lemma DERIV_cmult_right: hoelzl@56181: "(f has_field_derivative D) (at x within s) ==> ((\x. f x * c) has_field_derivative D * c) (at x within s)" haftmann@57514: using DERIV_cmult by (force simp add: ac_simps) lp15@55967: hoelzl@56181: lemma DERIV_cmult_Id [simp]: "(op * c has_field_derivative c) (at x within s)" hoelzl@51642: by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) hoelzl@51642: hoelzl@56181: lemma DERIV_cdivide: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" hoelzl@56181: using DERIV_cmult_right[of f D x s "1 / c"] by simp huffman@21164: huffman@21164: lemma DERIV_unique: hoelzl@51642: "DERIV f x :> D \ DERIV f x :> E \ D = E" hoelzl@56381: unfolding DERIV_def by (rule LIM_unique) huffman@21164: hoelzl@56381: lemma DERIV_setsum[derivative_intros]: hoelzl@56181: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ hoelzl@56181: ((\x. setsum (f x) S) has_field_derivative setsum (f' x) S) F" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_setsum]) hoelzl@56181: (auto simp: setsum_right_distrib mult_commute_abs dest: has_field_derivative_imp_has_derivative) hoelzl@51642: hoelzl@56381: lemma DERIV_inverse'[derivative_intros]: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ f x \ 0 \ hoelzl@56181: ((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_inverse]) hoelzl@56181: (auto dest: has_field_derivative_imp_has_derivative) hoelzl@51642: hoelzl@51642: text {* Power of @{text "-1"} *} hoelzl@51642: hoelzl@51642: lemma DERIV_inverse: hoelzl@56181: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" hoelzl@51642: by (drule DERIV_inverse' [OF DERIV_ident]) simp hoelzl@51642: hoelzl@51642: text {* Derivative of inverse *} hoelzl@51642: hoelzl@51642: lemma DERIV_inverse_fun: hoelzl@56181: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ hoelzl@56181: ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" haftmann@57514: by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) hoelzl@51642: hoelzl@51642: text {* Derivative of quotient *} hoelzl@51642: hoelzl@56381: lemma DERIV_divide[derivative_intros]: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ hoelzl@56181: (g has_field_derivative E) (at x within s) \ g x \ 0 \ hoelzl@56181: ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) hoelzl@56480: (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) hoelzl@51642: hoelzl@51642: lemma DERIV_quotient: hoelzl@56181: "(f has_field_derivative d) (at x within s) \ hoelzl@56181: (g has_field_derivative e) (at x within s)\ g x \ 0 \ hoelzl@56181: ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" haftmann@57512: by (drule (2) DERIV_divide) (simp add: mult.commute) hoelzl@51642: hoelzl@51642: lemma DERIV_power_Suc: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ hoelzl@56181: ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) hoelzl@56181: (auto simp: has_field_derivative_def) hoelzl@51642: hoelzl@56381: lemma DERIV_power[derivative_intros]: hoelzl@56181: "(f has_field_derivative D) (at x within s) \ hoelzl@56181: ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" hoelzl@56181: by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) hoelzl@56181: (auto simp: has_field_derivative_def) hoelzl@31880: hoelzl@56181: lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" hoelzl@51642: apply (cut_tac DERIV_power [OF DERIV_ident]) hoelzl@51642: apply (simp add: real_of_nat_def) hoelzl@51642: done hoelzl@51642: hoelzl@56181: lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ hoelzl@56181: ((\x. g (f x)) has_field_derivative E * D) (at x within s)" hoelzl@56181: using has_derivative_compose[of f "op * D" x s g "op * E"] hoelzl@56181: unfolding has_field_derivative_def mult_commute_abs ac_simps . hoelzl@51642: hoelzl@56181: corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ hoelzl@56181: ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" lp15@55967: by (rule DERIV_chain') lp15@55967: hoelzl@51642: text {* Standard version *} hoelzl@51642: hoelzl@56181: lemma DERIV_chain: hoelzl@56181: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ hoelzl@56181: (f o g has_field_derivative Da * Db) (at x within s)" haftmann@57512: by (drule (1) DERIV_chain', simp add: o_def mult.commute) hoelzl@51642: lp15@55967: lemma DERIV_image_chain: hoelzl@56181: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ hoelzl@56181: (f o g has_field_derivative Da * Db) (at x within s)" hoelzl@56181: using has_derivative_in_compose [of g "op * Db" x s f "op * Da "] hoelzl@56181: by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) lp15@55967: lp15@55967: (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lp15@55967: lemma DERIV_chain_s: lp15@55967: assumes "(\x. x \ s \ DERIV g x :> g'(x))" lp15@55967: and "DERIV f x :> f'" lp15@55967: and "f x \ s" lp15@55967: shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" haftmann@57512: by (metis (full_types) DERIV_chain' mult.commute assms) lp15@55967: lp15@55967: lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) lp15@55967: assumes "(\x. DERIV g x :> g'(x))" lp15@55967: and "DERIV f x :> f'" lp15@55967: shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" lp15@55967: by (metis UNIV_I DERIV_chain_s [of UNIV] assms) lp15@55967: hoelzl@51642: declare hoelzl@56381: DERIV_power[where 'a=real, unfolded real_of_nat_def[symmetric], derivative_intros] huffman@21164: huffman@21164: text{*Alternative definition for differentiability*} huffman@21164: huffman@21164: lemma DERIV_LIM_iff: huffman@31338: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows huffman@21784: "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) = huffman@21164: ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" huffman@21164: apply (rule iffI) huffman@21164: apply (drule_tac k="- a" in LIM_offset) haftmann@54230: apply simp huffman@21164: apply (drule_tac k="a" in LIM_offset) haftmann@57512: apply (simp add: add.commute) huffman@21164: done huffman@21164: hoelzl@51642: lemma DERIV_iff2: "(DERIV f x :> D) \ (\z. (f z - f x) / (z - x)) --x --> D" hoelzl@56381: by (simp add: DERIV_def DERIV_LIM_iff) huffman@21164: hoelzl@51642: lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ hoelzl@51642: DERIV f x :> u \ DERIV g y :> v" hoelzl@51642: unfolding DERIV_iff2 hoelzl@51642: proof (rule filterlim_cong) wenzelm@53374: assume *: "eventually (\x. f x = g x) (nhds x)" wenzelm@53374: moreover from * have "f x = g x" by (auto simp: eventually_nhds) hoelzl@51642: moreover assume "x = y" "u = v" hoelzl@51642: ultimately show "eventually (\xa. (f xa - f x) / (xa - x) = (g xa - g y) / (xa - y)) (at x)" hoelzl@51642: by (auto simp: eventually_at_filter elim: eventually_elim1) hoelzl@51642: qed simp_all huffman@21164: hoelzl@51642: lemma DERIV_shift: hoelzl@51642: "(DERIV f (x + z) :> y) \ (DERIV (\x. f (x + z)) x :> y)" hoelzl@56381: by (simp add: DERIV_def field_simps) huffman@21164: hoelzl@51642: lemma DERIV_mirror: hoelzl@51642: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x::real) :: real) x :> - y)" hoelzl@56479: by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right hoelzl@51642: tendsto_minus_cancel_left field_simps conj_commute) huffman@21164: huffman@29975: text {* Caratheodory formulation of derivative at a point *} huffman@21164: lp15@55970: lemma CARAT_DERIV: (*FIXME: SUPERSEDED BY THE ONE IN Deriv.thy. But still used by NSA/HDeriv.thy*) hoelzl@51642: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" huffman@21164: (is "?lhs = ?rhs") huffman@21164: proof huffman@21164: assume der: "DERIV f x :> l" huffman@21784: show "\g. (\z. f z - f x = g z * (z-x)) \ isCont g x \ g x = l" huffman@21164: proof (intro exI conjI) huffman@21784: let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" nipkow@23413: show "\z. f z - f x = ?g z * (z-x)" by simp huffman@21164: show "isCont ?g x" using der hoelzl@56381: by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) huffman@21164: show "?g x = l" by simp huffman@21164: qed huffman@21164: next huffman@21164: assume "?rhs" huffman@21164: then obtain g where huffman@21784: "(\z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast huffman@21164: thus "(DERIV f x :> l)" hoelzl@56381: by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) huffman@21164: qed huffman@21164: wenzelm@31899: text {* wenzelm@31899: Let's do the standard proof, though theorem wenzelm@31899: @{text "LIM_mult2"} follows from a NS proof wenzelm@31899: *} huffman@21164: huffman@29975: subsection {* Local extrema *} huffman@29975: huffman@21164: text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} huffman@21164: paulson@33654: lemma DERIV_pos_inc_right: huffman@21164: fixes f :: "real => real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and l: "0 < l" huffman@21164: shows "\d > 0. \h > 0. h < d --> f(x) < f(x + h)" huffman@21164: proof - huffman@21164: from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] huffman@21164: have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l)" haftmann@54230: by simp huffman@21164: then obtain s huffman@21164: where s: "0 < s" huffman@21164: and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < l" huffman@21164: by auto huffman@21164: thus ?thesis huffman@21164: proof (intro exI conjI strip) huffman@23441: show "0 real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and l: "l < 0" huffman@21164: shows "\d > 0. \h > 0. h < d --> f(x) < f(x-h)" huffman@21164: proof - huffman@21164: from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] huffman@21164: have "\s > 0. (\z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l)" haftmann@54230: by simp huffman@21164: then obtain s huffman@21164: where s: "0 < s" huffman@21164: and all: "!!z. z \ 0 \ \z\ < s \ \(f(x+z) - f x) / z - l\ < -l" huffman@21164: by auto huffman@21164: thus ?thesis huffman@21164: proof (intro exI conjI strip) huffman@23441: show "0 real" paulson@33654: shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d --> f(x - h) < f(x)" hoelzl@56181: apply (rule DERIV_neg_dec_left [of "%x. - f x" "-l" x, simplified]) hoelzl@41368: apply (auto simp add: DERIV_minus) paulson@33654: done paulson@33654: paulson@33654: lemma DERIV_neg_dec_right: paulson@33654: fixes f :: "real => real" paulson@33654: shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d --> f(x) > f(x + h)" hoelzl@56181: apply (rule DERIV_pos_inc_right [of "%x. - f x" "-l" x, simplified]) hoelzl@41368: apply (auto simp add: DERIV_minus) paulson@33654: done paulson@33654: huffman@21164: lemma DERIV_local_max: huffman@21164: fixes f :: "real => real" huffman@21164: assumes der: "DERIV f x :> l" huffman@21164: and d: "0 < d" huffman@21164: and le: "\y. \x-y\ < d --> f(y) \ f(x)" huffman@21164: shows "l = 0" huffman@21164: proof (cases rule: linorder_cases [of l 0]) huffman@23441: case equal thus ?thesis . huffman@21164: next huffman@21164: case less paulson@33654: from DERIV_neg_dec_left [OF der less] huffman@21164: obtain d' where d': "0 < d'" huffman@21164: and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast huffman@21164: from real_lbound_gt_zero [OF d d'] huffman@21164: obtain e where "0 < e \ e < d \ e < d'" .. huffman@21164: with lt le [THEN spec [where x="x-e"]] huffman@21164: show ?thesis by (auto simp add: abs_if) huffman@21164: next huffman@21164: case greater paulson@33654: from DERIV_pos_inc_right [OF der greater] huffman@21164: obtain d' where d': "0 < d'" huffman@21164: and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast huffman@21164: from real_lbound_gt_zero [OF d d'] huffman@21164: obtain e where "0 < e \ e < d \ e < d'" .. huffman@21164: with lt le [THEN spec [where x="x+e"]] huffman@21164: show ?thesis by (auto simp add: abs_if) huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: text{*Similar theorem for a local minimum*} huffman@21164: lemma DERIV_local_min: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) \ f(y) |] ==> l = 0" huffman@21164: by (drule DERIV_minus [THEN DERIV_local_max], auto) huffman@21164: huffman@21164: huffman@21164: text{*In particular, if a function is locally flat*} huffman@21164: lemma DERIV_local_const: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| DERIV f x :> l; 0 < d; \y. \x-y\ < d --> f(x) = f(y) |] ==> l = 0" huffman@21164: by (auto dest!: DERIV_local_max) huffman@21164: huffman@29975: huffman@29975: subsection {* Rolle's Theorem *} huffman@29975: huffman@21164: text{*Lemma about introducing open ball in open interval*} huffman@21164: lemma lemma_interval_lt: huffman@21164: "[| a < x; x < b |] huffman@21164: ==> \d::real. 0 < d & (\y. \x-y\ < d --> a < y & y < b)" chaieb@27668: huffman@22998: apply (simp add: abs_less_iff) huffman@21164: apply (insert linorder_linear [of "x-a" "b-x"], safe) huffman@21164: apply (rule_tac x = "x-a" in exI) huffman@21164: apply (rule_tac [2] x = "b-x" in exI, auto) huffman@21164: done huffman@21164: huffman@21164: lemma lemma_interval: "[| a < x; x < b |] ==> huffman@21164: \d::real. 0 < d & (\y. \x-y\ < d --> a \ y & y \ b)" huffman@21164: apply (drule lemma_interval_lt, auto) huffman@44921: apply force huffman@21164: done huffman@21164: huffman@21164: text{*Rolle's Theorem. huffman@21164: If @{term f} is defined and continuous on the closed interval huffman@21164: @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"}, huffman@21164: and @{term "f(a) = f(b)"}, huffman@21164: then there exists @{text "x0 \ (a,b)"} such that @{term "f'(x0) = 0"}*} huffman@21164: theorem Rolle: huffman@21164: assumes lt: "a < b" huffman@21164: and eq: "f(a) = f(b)" huffman@21164: and con: "\x. a \ x & x \ b --> isCont f x" hoelzl@56181: and dif [rule_format]: "\x. a < x & x < b --> f differentiable (at x)" huffman@21784: shows "\z::real. a < z & z < b & DERIV f z :> 0" huffman@21164: proof - huffman@21164: have le: "a \ b" using lt by simp huffman@21164: from isCont_eq_Ub [OF le con] huffman@21164: obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" huffman@21164: and alex: "a \ x" and xleb: "x \ b" huffman@21164: by blast huffman@21164: from isCont_eq_Lb [OF le con] huffman@21164: obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" huffman@21164: and alex': "a \ x'" and x'leb: "x' \ b" huffman@21164: by blast huffman@21164: show ?thesis huffman@21164: proof cases huffman@21164: assume axb: "a < x & x < b" huffman@21164: --{*@{term f} attains its maximum within the interval*} chaieb@27668: hence ax: "ay. \x-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: hence bound': "\y. \x-y\ < d \ f y \ f x" using x_max huffman@21164: by blast huffman@21164: from differentiableD [OF dif [OF axb]] huffman@21164: obtain l where der: "DERIV f x :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_max [OF der d bound']) huffman@21164: --{*the derivative at a local maximum is zero*} huffman@21164: thus ?thesis using ax xb der by auto huffman@21164: next huffman@21164: assume notaxb: "~ (a < x & x < b)" huffman@21164: hence xeqab: "x=a | x=b" using alex xleb by arith huffman@21164: hence fb_eq_fx: "f b = f x" by (auto simp add: eq) huffman@21164: show ?thesis huffman@21164: proof cases huffman@21164: assume ax'b: "a < x' & x' < b" huffman@21164: --{*@{term f} attains its minimum within the interval*} chaieb@27668: hence ax': "ay. \x'-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: hence bound': "\y. \x'-y\ < d \ f x' \ f y" using x'_min huffman@21164: by blast huffman@21164: from differentiableD [OF dif [OF ax'b]] huffman@21164: obtain l where der: "DERIV f x' :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_min [OF der d bound']) huffman@21164: --{*the derivative at a local minimum is zero*} huffman@21164: thus ?thesis using ax' x'b der by auto huffman@21164: next huffman@21164: assume notax'b: "~ (a < x' & x' < b)" huffman@21164: --{*@{term f} is constant througout the interval*} huffman@21164: hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith huffman@21164: hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) huffman@21164: from dense [OF lt] huffman@21164: obtain r where ar: "a < r" and rb: "r < b" by blast huffman@21164: from lemma_interval [OF ar rb] huffman@21164: obtain d where d: "0y. \r-y\ < d \ a \ y \ y \ b" huffman@21164: by blast huffman@21164: have eq_fb: "\z. a \ z --> z \ b --> f z = f b" huffman@21164: proof (clarify) huffman@21164: fix z::real huffman@21164: assume az: "a \ z" and zb: "z \ b" huffman@21164: show "f z = f b" huffman@21164: proof (rule order_antisym) huffman@21164: show "f z \ f b" by (simp add: fb_eq_fx x_max az zb) huffman@21164: show "f b \ f z" by (simp add: fb_eq_fx' x'_min az zb) huffman@21164: qed huffman@21164: qed huffman@21164: have bound': "\y. \r-y\ < d \ f r = f y" huffman@21164: proof (intro strip) huffman@21164: fix y::real huffman@21164: assume lt: "\r-y\ < d" huffman@21164: hence "f y = f b" by (simp add: eq_fb bound) huffman@21164: thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) huffman@21164: qed huffman@21164: from differentiableD [OF dif [OF conjI [OF ar rb]]] huffman@21164: obtain l where der: "DERIV f r :> l" .. huffman@21164: have "l=0" by (rule DERIV_local_const [OF der d bound']) huffman@21164: --{*the derivative of a constant function is zero*} huffman@21164: thus ?thesis using ar rb der by auto huffman@21164: qed huffman@21164: qed huffman@21164: qed huffman@21164: huffman@21164: huffman@21164: subsection{*Mean Value Theorem*} huffman@21164: huffman@21164: lemma lemma_MVT: huffman@21164: "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" hoelzl@51481: by (cases "a = b") (simp_all add: field_simps) huffman@21164: huffman@21164: theorem MVT: huffman@21164: assumes lt: "a < b" huffman@21164: and con: "\x. a \ x & x \ b --> isCont f x" hoelzl@56181: and dif [rule_format]: "\x. a < x & x < b --> f differentiable (at x)" huffman@21784: shows "\l z::real. a < z & z < b & DERIV f z :> l & huffman@21164: (f(b) - f(a) = (b-a) * l)" huffman@21164: proof - huffman@21164: let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" huffman@44233: have contF: "\x. a \ x \ x \ b \ isCont ?F x" hoelzl@56371: using con by (fast intro: continuous_intros) hoelzl@56181: have difF: "\x. a < x \ x < b \ ?F differentiable (at x)" huffman@21164: proof (clarify) huffman@21164: fix x::real huffman@21164: assume ax: "a < x" and xb: "x < b" huffman@21164: from differentiableD [OF dif [OF conjI [OF ax xb]]] huffman@21164: obtain l where der: "DERIV f x :> l" .. hoelzl@56181: show "?F differentiable (at x)" huffman@21164: by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], huffman@21164: blast intro: DERIV_diff DERIV_cmult_Id der) huffman@21164: qed huffman@21164: from Rolle [where f = ?F, OF lt lemma_MVT contF difF] huffman@21164: obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" huffman@21164: by blast huffman@21164: have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" huffman@21164: by (rule DERIV_cmult_Id) huffman@21164: hence derF: "DERIV (\x. ?F x + (f b - f a) / (b - a) * x) z huffman@21164: :> 0 + (f b - f a) / (b - a)" huffman@21164: by (rule DERIV_add [OF der]) huffman@21164: show ?thesis huffman@21164: proof (intro exI conjI) huffman@23441: show "a < z" using az . huffman@23441: show "z < b" using zb . huffman@21164: show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) huffman@21164: show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp huffman@21164: qed huffman@21164: qed huffman@21164: hoelzl@29803: lemma MVT2: hoelzl@29803: "[| a < b; \x. a \ x & x \ b --> DERIV f x :> f'(x) |] hoelzl@29803: ==> \z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))" hoelzl@29803: apply (drule MVT) hoelzl@29803: apply (blast intro: DERIV_isCont) hoelzl@56181: apply (force dest: order_less_imp_le simp add: real_differentiable_def) hoelzl@29803: apply (blast dest: DERIV_unique order_less_imp_le) hoelzl@29803: done hoelzl@29803: huffman@21164: huffman@21164: text{*A function is constant if its derivative is 0 over an interval.*} huffman@21164: huffman@21164: lemma DERIV_isconst_end: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0 |] huffman@21164: ==> f b = f a" huffman@21164: apply (drule MVT, assumption) huffman@21164: apply (blast intro: differentiableI) huffman@21164: apply (auto dest!: DERIV_unique simp add: diff_eq_eq) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_isconst1: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0 |] huffman@21164: ==> \x. a \ x & x \ b --> f x = f a" huffman@21164: apply safe huffman@21164: apply (drule_tac x = a in order_le_imp_less_or_eq, safe) huffman@21164: apply (drule_tac b = x in DERIV_isconst_end, auto) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_isconst2: huffman@21164: fixes f :: "real => real" huffman@21164: shows "[| a < b; huffman@21164: \x. a \ x & x \ b --> isCont f x; huffman@21164: \x. a < x & x < b --> DERIV f x :> 0; huffman@21164: a \ x; x \ b |] huffman@21164: ==> f x = f a" huffman@21164: apply (blast dest: DERIV_isconst1) huffman@21164: done huffman@21164: hoelzl@29803: lemma DERIV_isconst3: fixes a b x y :: real hoelzl@29803: assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" hoelzl@29803: assumes derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" hoelzl@29803: shows "f x = f y" hoelzl@29803: proof (cases "x = y") hoelzl@29803: case False hoelzl@29803: let ?a = "min x y" hoelzl@29803: let ?b = "max x y" hoelzl@29803: hoelzl@29803: have "\z. ?a \ z \ z \ ?b \ DERIV f z :> 0" hoelzl@29803: proof (rule allI, rule impI) hoelzl@29803: fix z :: real assume "?a \ z \ z \ ?b" hoelzl@29803: hence "a < z" and "z < b" using `x \ {a <..< b}` and `y \ {a <..< b}` by auto hoelzl@29803: hence "z \ {a<.. 0" by (rule derivable) hoelzl@29803: qed hoelzl@29803: hence isCont: "\z. ?a \ z \ z \ ?b \ isCont f z" hoelzl@29803: and DERIV: "\z. ?a < z \ z < ?b \ DERIV f z :> 0" using DERIV_isCont by auto hoelzl@29803: hoelzl@29803: have "?a < ?b" using `x \ y` by auto hoelzl@29803: from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] hoelzl@29803: show ?thesis by auto hoelzl@29803: qed auto hoelzl@29803: huffman@21164: lemma DERIV_isconst_all: huffman@21164: fixes f :: "real => real" huffman@21164: shows "\x. DERIV f x :> 0 ==> f(x) = f(y)" huffman@21164: apply (rule linorder_cases [of x y]) huffman@21164: apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_const_ratio_const: huffman@21784: fixes f :: "real => real" huffman@21784: shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" huffman@21164: apply (rule linorder_cases [of a b], auto) huffman@21164: apply (drule_tac [!] f = f in MVT) hoelzl@56181: apply (auto dest: DERIV_isCont DERIV_unique simp add: real_differentiable_def) haftmann@54230: apply (auto dest: DERIV_unique simp add: ring_distribs) huffman@21164: done huffman@21164: huffman@21164: lemma DERIV_const_ratio_const2: huffman@21784: fixes f :: "real => real" huffman@21784: shows "[|a \ b; \x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" lp15@56217: apply (rule_tac c1 = "b-a" in mult_right_cancel [THEN iffD1]) haftmann@57512: apply (auto dest!: DERIV_const_ratio_const simp add: mult.assoc) huffman@21164: done huffman@21164: huffman@21164: lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" huffman@21164: by (simp) huffman@21164: huffman@21164: lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" huffman@21164: by (simp) huffman@21164: huffman@21164: text{*Gallileo's "trick": average velocity = av. of end velocities*} huffman@21164: huffman@21164: lemma DERIV_const_average: huffman@21164: fixes v :: "real => real" huffman@21164: assumes neq: "a \ (b::real)" huffman@21164: and der: "\x. DERIV v x :> k" huffman@21164: shows "v ((a + b)/2) = (v a + v b)/2" huffman@21164: proof (cases rule: linorder_cases [of a b]) huffman@21164: case equal with neq show ?thesis by simp huffman@21164: next huffman@21164: case less huffman@21164: have "(v b - v a) / (b - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF neq der]) huffman@21164: hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp huffman@21164: moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) huffman@21164: ultimately show ?thesis using neq by force huffman@21164: next huffman@21164: case greater huffman@21164: have "(v b - v a) / (b - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF neq der]) huffman@21164: hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp huffman@21164: moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" huffman@21164: by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) haftmann@57512: ultimately show ?thesis using neq by (force simp add: add.commute) huffman@21164: qed huffman@21164: paulson@33654: (* A function with positive derivative is increasing. paulson@33654: A simple proof using the MVT, by Jeremy Avigad. And variants. paulson@33654: *) lp15@56261: lemma DERIV_pos_imp_increasing_open: paulson@33654: fixes a::real and b::real and f::"real => real" lp15@56261: assumes "a < b" and "\x. a < x \ x < b \ (EX y. DERIV f x :> y & y > 0)" lp15@56261: and con: "\x. a \ x \ x \ b \ isCont f x" paulson@33654: shows "f a < f b" paulson@33654: proof (rule ccontr) wenzelm@41550: assume f: "~ f a < f b" wenzelm@33690: have "EX l z. a < z & z < b & DERIV f z :> l paulson@33654: & f b - f a = (b - a) * l" wenzelm@33690: apply (rule MVT) lp15@56261: using assms Deriv.differentiableI lp15@56261: apply force+ wenzelm@33690: done wenzelm@41550: then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" paulson@33654: and "f b - f a = (b - a) * l" paulson@33654: by auto wenzelm@41550: with assms f have "~(l > 0)" huffman@36777: by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) wenzelm@41550: with assms z show False lp15@56261: by (metis DERIV_unique) paulson@33654: qed paulson@33654: lp15@56261: lemma DERIV_pos_imp_increasing: lp15@56261: fixes a::real and b::real and f::"real => real" lp15@56261: assumes "a < b" and "\x. a \ x & x \ b --> (EX y. DERIV f x :> y & y > 0)" lp15@56261: shows "f a < f b" lp15@56261: by (metis DERIV_pos_imp_increasing_open [of a b f] assms DERIV_continuous less_imp_le) lp15@56261: noschinl@45791: lemma DERIV_nonneg_imp_nondecreasing: paulson@33654: fixes a::real and b::real and f::"real => real" paulson@33654: assumes "a \ b" and paulson@33654: "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" paulson@33654: shows "f a \ f b" paulson@33654: proof (rule ccontr, cases "a = b") wenzelm@41550: assume "~ f a \ f b" and "a = b" wenzelm@41550: then show False by auto haftmann@37891: next haftmann@37891: assume A: "~ f a \ f b" haftmann@37891: assume B: "a ~= b" paulson@33654: with assms have "EX l z. a < z & z < b & DERIV f z :> l paulson@33654: & f b - f a = (b - a) * l" wenzelm@33690: apply - wenzelm@33690: apply (rule MVT) wenzelm@33690: apply auto wenzelm@33690: apply (metis DERIV_isCont) huffman@36777: apply (metis differentiableI less_le) paulson@33654: done wenzelm@41550: then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" haftmann@37891: and C: "f b - f a = (b - a) * l" paulson@33654: by auto haftmann@37891: with A have "a < b" "f b < f a" by auto haftmann@37891: with C have "\ l \ 0" by (auto simp add: not_le algebra_simps) huffman@45051: (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) wenzelm@41550: with assms z show False paulson@33654: by (metis DERIV_unique order_less_imp_le) paulson@33654: qed paulson@33654: lp15@56261: lemma DERIV_neg_imp_decreasing_open: lp15@56261: fixes a::real and b::real and f::"real => real" lp15@56261: assumes "a < b" and "\x. a < x \ x < b \ (EX y. DERIV f x :> y & y < 0)" lp15@56261: and con: "\x. a \ x \ x \ b \ isCont f x" lp15@56261: shows "f a > f b" lp15@56261: proof - lp15@56261: have "(%x. -f x) a < (%x. -f x) b" lp15@56261: apply (rule DERIV_pos_imp_increasing_open [of a b "%x. -f x"]) lp15@56261: using assms lp15@56261: apply auto lp15@56261: apply (metis field_differentiable_minus neg_0_less_iff_less) lp15@56261: done lp15@56261: thus ?thesis lp15@56261: by simp lp15@56261: qed lp15@56261: paulson@33654: lemma DERIV_neg_imp_decreasing: paulson@33654: fixes a::real and b::real and f::"real => real" paulson@33654: assumes "a < b" and paulson@33654: "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y < 0)" paulson@33654: shows "f a > f b" lp15@56261: by (metis DERIV_neg_imp_decreasing_open [of a b f] assms DERIV_continuous less_imp_le) paulson@33654: paulson@33654: lemma DERIV_nonpos_imp_nonincreasing: paulson@33654: fixes a::real and b::real and f::"real => real" paulson@33654: assumes "a \ b" and paulson@33654: "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" paulson@33654: shows "f a \ f b" paulson@33654: proof - paulson@33654: have "(%x. -f x) a \ (%x. -f x) b" noschinl@45791: apply (rule DERIV_nonneg_imp_nondecreasing [of a b "%x. -f x"]) wenzelm@33690: using assms wenzelm@33690: apply auto paulson@33654: apply (metis DERIV_minus neg_0_le_iff_le) paulson@33654: done paulson@33654: thus ?thesis paulson@33654: by simp paulson@33654: qed huffman@21164: lp15@56289: lemma DERIV_pos_imp_increasing_at_bot: lp15@56289: fixes f :: "real => real" lp15@56289: assumes "\x. x \ b \ (EX y. DERIV f x :> y & y > 0)" lp15@56289: and lim: "(f ---> flim) at_bot" lp15@56289: shows "flim < f b" lp15@56289: proof - lp15@56289: have "flim \ f (b - 1)" lp15@56289: apply (rule tendsto_ge_const [OF _ lim]) lp15@56289: apply (auto simp: trivial_limit_at_bot_linorder eventually_at_bot_linorder) lp15@56289: apply (rule_tac x="b - 2" in exI) lp15@56289: apply (force intro: order.strict_implies_order DERIV_pos_imp_increasing [where f=f] assms) lp15@56289: done lp15@56289: also have "... < f b" lp15@56289: by (force intro: DERIV_pos_imp_increasing [where f=f] assms) lp15@56289: finally show ?thesis . lp15@56289: qed lp15@56289: lp15@56289: lemma DERIV_neg_imp_decreasing_at_top: lp15@56289: fixes f :: "real => real" lp15@56289: assumes der: "\x. x \ b \ (EX y. DERIV f x :> y & y < 0)" lp15@56289: and lim: "(f ---> flim) at_top" lp15@56289: shows "flim < f b" lp15@56289: apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) lp15@56289: apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) lp15@56289: apply (metis filterlim_at_top_mirror lim) lp15@56289: done lp15@56289: huffman@23041: text {* Derivative of inverse function *} huffman@23041: huffman@23041: lemma DERIV_inverse_function: huffman@23041: fixes f g :: "real \ real" huffman@23041: assumes der: "DERIV f (g x) :> D" huffman@23041: assumes neq: "D \ 0" huffman@23044: assumes a: "a < x" and b: "x < b" huffman@23044: assumes inj: "\y. a < y \ y < b \ f (g y) = y" huffman@23041: assumes cont: "isCont g x" huffman@23041: shows "DERIV g x :> inverse D" huffman@23041: unfolding DERIV_iff2 huffman@23044: proof (rule LIM_equal2) huffman@23044: show "0 < min (x - a) (b - x)" chaieb@27668: using a b by arith huffman@23044: next huffman@23041: fix y huffman@23044: assume "norm (y - x) < min (x - a) (b - x)" chaieb@27668: hence "a < y" and "y < b" huffman@23044: by (simp_all add: abs_less_iff) huffman@23041: thus "(g y - g x) / (y - x) = huffman@23041: inverse ((f (g y) - x) / (g y - g x))" huffman@23041: by (simp add: inj) huffman@23041: next huffman@23041: have "(\z. (f z - f (g x)) / (z - g x)) -- g x --> D" huffman@23041: by (rule der [unfolded DERIV_iff2]) huffman@23041: hence 1: "(\z. (f z - x) / (z - g x)) -- g x --> D" huffman@23044: using inj a b by simp huffman@23041: have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" wenzelm@56219: proof (rule exI, safe) huffman@23044: show "0 < min (x - a) (b - x)" huffman@23044: using a b by simp huffman@23041: next huffman@23041: fix y huffman@23044: assume "norm (y - x) < min (x - a) (b - x)" huffman@23044: hence y: "a < y" "y < b" huffman@23044: by (simp_all add: abs_less_iff) huffman@23041: assume "g y = g x" huffman@23041: hence "f (g y) = f (g x)" by simp huffman@23044: hence "y = x" using inj y a b by simp huffman@23041: also assume "y \ x" huffman@23041: finally show False by simp huffman@23041: qed huffman@23041: have "(\y. (f (g y) - x) / (g y - g x)) -- x --> D" huffman@23041: using cont 1 2 by (rule isCont_LIM_compose2) huffman@23041: thus "(\y. inverse ((f (g y) - x) / (g y - g x))) huffman@23041: -- x --> inverse D" huffman@44568: using neq by (rule tendsto_inverse) huffman@23041: qed huffman@23041: huffman@29975: subsection {* Generalized Mean Value Theorem *} huffman@29975: huffman@21164: theorem GMVT: huffman@21784: fixes a b :: real huffman@21164: assumes alb: "a < b" wenzelm@41550: and fc: "\x. a \ x \ x \ b \ isCont f x" hoelzl@56181: and fd: "\x. a < x \ x < b \ f differentiable (at x)" wenzelm@41550: and gc: "\x. a \ x \ x \ b \ isCont g x" hoelzl@56181: and gd: "\x. a < x \ x < b \ g differentiable (at x)" wenzelm@53381: shows "\g'c f'c c. wenzelm@53381: DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ ((f b - f a) * g'c) = ((g b - g a) * f'c)" huffman@21164: proof - huffman@21164: let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" wenzelm@41550: from assms have "a < b" by simp huffman@21164: moreover have "\x. a \ x \ x \ b \ isCont ?h x" huffman@44233: using fc gc by simp hoelzl@56181: moreover have "\x. a < x \ x < b \ ?h differentiable (at x)" huffman@44233: using fd gd by simp huffman@21164: ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) huffman@21164: then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. huffman@21164: then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. huffman@21164: huffman@21164: from cdef have cint: "a < c \ c < b" by auto hoelzl@56181: with gd have "g differentiable (at c)" by simp huffman@21164: hence "\D. DERIV g c :> D" by (rule differentiableD) huffman@21164: then obtain g'c where g'cdef: "DERIV g c :> g'c" .. huffman@21164: huffman@21164: from cdef have "a < c \ c < b" by auto hoelzl@56181: with fd have "f differentiable (at c)" by simp huffman@21164: hence "\D. DERIV f c :> D" by (rule differentiableD) huffman@21164: then obtain f'c where f'cdef: "DERIV f c :> f'c" .. huffman@21164: huffman@21164: from cdef have "DERIV ?h c :> l" by auto hoelzl@41368: moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" hoelzl@56381: using g'cdef f'cdef by (auto intro!: derivative_eq_intros) huffman@21164: ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) huffman@21164: huffman@21164: { huffman@21164: from cdef have "?h b - ?h a = (b - a) * l" by auto wenzelm@53374: also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp huffman@21164: finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp huffman@21164: } huffman@21164: moreover huffman@21164: { huffman@21164: have "?h b - ?h a = huffman@21164: ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - huffman@21164: ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" nipkow@29667: by (simp add: algebra_simps) huffman@21164: hence "?h b - ?h a = 0" by auto huffman@21164: } huffman@21164: ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto huffman@21164: with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp huffman@21164: hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp haftmann@57514: hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) huffman@21164: huffman@21164: with g'cdef f'cdef cint show ?thesis by auto huffman@21164: qed huffman@21164: hoelzl@50327: lemma GMVT': hoelzl@50327: fixes f g :: "real \ real" hoelzl@50327: assumes "a < b" hoelzl@50327: assumes isCont_f: "\z. a \ z \ z \ b \ isCont f z" hoelzl@50327: assumes isCont_g: "\z. a \ z \ z \ b \ isCont g z" hoelzl@50327: assumes DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" hoelzl@50327: assumes DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" hoelzl@50327: shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" hoelzl@50327: proof - hoelzl@50327: have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ hoelzl@50327: a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" hoelzl@56181: using assms by (intro GMVT) (force simp: real_differentiable_def)+ hoelzl@50327: then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" hoelzl@50327: using DERIV_f DERIV_g by (force dest: DERIV_unique) hoelzl@50327: then show ?thesis hoelzl@50327: by auto hoelzl@50327: qed hoelzl@50327: hoelzl@51529: hoelzl@51529: subsection {* L'Hopitals rule *} hoelzl@51529: hoelzl@51641: lemma isCont_If_ge: hoelzl@51641: fixes a :: "'a :: linorder_topology" hoelzl@51641: shows "continuous (at_left a) g \ (f ---> g a) (at_right a) \ isCont (\x. if x \ a then g x else f x) a" hoelzl@51641: unfolding isCont_def continuous_within hoelzl@51641: apply (intro filterlim_split_at) hoelzl@51641: apply (subst filterlim_cong[OF refl refl, where g=g]) hoelzl@51641: apply (simp_all add: eventually_at_filter less_le) hoelzl@51641: apply (subst filterlim_cong[OF refl refl, where g=f]) hoelzl@51641: apply (simp_all add: eventually_at_filter less_le) hoelzl@51641: done hoelzl@51641: hoelzl@50327: lemma lhopital_right_0: hoelzl@50329: fixes f0 g0 :: "real \ real" hoelzl@50329: assumes f_0: "(f0 ---> 0) (at_right 0)" hoelzl@50329: assumes g_0: "(g0 ---> 0) (at_right 0)" hoelzl@50327: assumes ev: hoelzl@50329: "eventually (\x. g0 x \ 0) (at_right 0)" hoelzl@50327: "eventually (\x. g' x \ 0) (at_right 0)" hoelzl@50329: "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" hoelzl@50329: "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" hoelzl@50327: assumes lim: "((\ x. (f' x / g' x)) ---> x) (at_right 0)" hoelzl@50329: shows "((\ x. f0 x / g0 x) ---> x) (at_right 0)" hoelzl@50327: proof - hoelzl@50329: def f \ "\x. if x \ 0 then 0 else f0 x" hoelzl@50329: then have "f 0 = 0" by simp hoelzl@50329: hoelzl@50329: def g \ "\x. if x \ 0 then 0 else g0 x" hoelzl@50329: then have "g 0 = 0" by simp hoelzl@50329: hoelzl@50329: have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ hoelzl@50329: DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" hoelzl@50329: using ev by eventually_elim auto hoelzl@50329: then obtain a where [arith]: "0 < a" hoelzl@50329: and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" hoelzl@50327: and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" hoelzl@50329: and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" hoelzl@50329: and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" wenzelm@56219: unfolding eventually_at by (auto simp: dist_real_def) hoelzl@50327: hoelzl@50329: have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" hoelzl@50329: using g0_neq_0 by (simp add: g_def) hoelzl@50329: hoelzl@50329: { fix x assume x: "0 < x" "x < a" then have "DERIV f x :> (f' x)" hoelzl@50329: by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) hoelzl@50329: (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) } hoelzl@50329: note f = this hoelzl@50329: hoelzl@50329: { fix x assume x: "0 < x" "x < a" then have "DERIV g x :> (g' x)" hoelzl@50329: by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) hoelzl@50329: (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) } hoelzl@50329: note g = this hoelzl@50329: hoelzl@50329: have "isCont f 0" hoelzl@51641: unfolding f_def by (intro isCont_If_ge f_0 continuous_const) hoelzl@51641: hoelzl@50329: have "isCont g 0" hoelzl@51641: unfolding g_def by (intro isCont_If_ge g_0 continuous_const) hoelzl@50329: hoelzl@50327: have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" hoelzl@50327: proof (rule bchoice, rule) hoelzl@50327: fix x assume "x \ {0 <..< a}" hoelzl@50327: then have x[arith]: "0 < x" "x < a" by auto hoelzl@50327: with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" hoelzl@50327: by auto hoelzl@50328: have "\x. 0 \ x \ x < a \ isCont f x" hoelzl@50328: using `isCont f 0` f by (auto intro: DERIV_isCont simp: le_less) hoelzl@50328: moreover have "\x. 0 \ x \ x < a \ isCont g x" hoelzl@50328: using `isCont g 0` g by (auto intro: DERIV_isCont simp: le_less) hoelzl@50328: ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" hoelzl@50328: using f g `x < a` by (intro GMVT') auto wenzelm@53374: then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" wenzelm@53374: by blast hoelzl@50327: moreover wenzelm@53374: from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" hoelzl@50327: by (simp add: field_simps) hoelzl@50327: ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" hoelzl@50327: using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c]) hoelzl@50327: qed wenzelm@53381: then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. hoelzl@50327: then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" hoelzl@51641: unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) hoelzl@50327: moreover hoelzl@50327: from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" hoelzl@50327: by eventually_elim auto hoelzl@50327: then have "((\x. norm (\ x)) ---> 0) (at_right 0)" hoelzl@50327: by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) hoelzl@51641: (auto intro: tendsto_const tendsto_ident_at) hoelzl@50327: then have "(\ ---> 0) (at_right 0)" hoelzl@50327: by (rule tendsto_norm_zero_cancel) hoelzl@50327: with \ have "filterlim \ (at_right 0) (at_right 0)" hoelzl@51641: by (auto elim!: eventually_elim1 simp: filterlim_at) hoelzl@50327: from this lim have "((\t. f' (\ t) / g' (\ t)) ---> x) (at_right 0)" hoelzl@50327: by (rule_tac filterlim_compose[of _ _ _ \]) hoelzl@50329: ultimately have "((\t. f t / g t) ---> x) (at_right 0)" (is ?P) hoelzl@50328: by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) hoelzl@50328: (auto elim: eventually_elim1) hoelzl@50329: also have "?P \ ?thesis" hoelzl@51641: by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) hoelzl@50329: finally show ?thesis . hoelzl@50327: qed hoelzl@50327: hoelzl@50330: lemma lhopital_right: hoelzl@50330: "((f::real \ real) ---> 0) (at_right x) \ (g ---> 0) (at_right x) \ hoelzl@50330: eventually (\x. g x \ 0) (at_right x) \ hoelzl@50330: eventually (\x. g' x \ 0) (at_right x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at_right x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at_right x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at_right x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at_right x)" hoelzl@50330: unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift hoelzl@50330: by (rule lhopital_right_0) hoelzl@50330: hoelzl@50330: lemma lhopital_left: hoelzl@50330: "((f::real \ real) ---> 0) (at_left x) \ (g ---> 0) (at_left x) \ hoelzl@50330: eventually (\x. g x \ 0) (at_left x) \ hoelzl@50330: eventually (\x. g' x \ 0) (at_left x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at_left x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at_left x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at_left x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at_left x)" hoelzl@50330: unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror hoelzl@56479: by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) hoelzl@50330: hoelzl@50330: lemma lhopital: hoelzl@50330: "((f::real \ real) ---> 0) (at x) \ (g ---> 0) (at x) \ hoelzl@50330: eventually (\x. g x \ 0) (at x) \ hoelzl@50330: eventually (\x. g' x \ 0) (at x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at x)" hoelzl@50330: unfolding eventually_at_split filterlim_at_split hoelzl@50330: by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) hoelzl@50330: hoelzl@50327: lemma lhopital_right_0_at_top: hoelzl@50327: fixes f g :: "real \ real" hoelzl@50327: assumes g_0: "LIM x at_right 0. g x :> at_top" hoelzl@50327: assumes ev: hoelzl@50327: "eventually (\x. g' x \ 0) (at_right 0)" hoelzl@50327: "eventually (\x. DERIV f x :> f' x) (at_right 0)" hoelzl@50327: "eventually (\x. DERIV g x :> g' x) (at_right 0)" hoelzl@50327: assumes lim: "((\ x. (f' x / g' x)) ---> x) (at_right 0)" hoelzl@50327: shows "((\ x. f x / g x) ---> x) (at_right 0)" hoelzl@50327: unfolding tendsto_iff hoelzl@50327: proof safe hoelzl@50327: fix e :: real assume "0 < e" hoelzl@50327: hoelzl@50327: with lim[unfolded tendsto_iff, rule_format, of "e / 4"] hoelzl@50327: have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp hoelzl@50327: from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] hoelzl@50327: obtain a where [arith]: "0 < a" hoelzl@50327: and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" hoelzl@50327: and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" hoelzl@50327: and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" hoelzl@50327: and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" hoelzl@51641: unfolding eventually_at_le by (auto simp: dist_real_def) hoelzl@51641: hoelzl@50327: hoelzl@50327: from Df have hoelzl@50328: "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" hoelzl@51641: unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) hoelzl@50327: hoelzl@50327: moreover hoelzl@50328: have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" hoelzl@50346: using g_0 by (auto elim: eventually_elim1 simp: filterlim_at_top_dense) hoelzl@50327: hoelzl@50327: moreover hoelzl@50327: have inv_g: "((\x. inverse (g x)) ---> 0) (at_right 0)" hoelzl@50327: using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] hoelzl@50327: by (rule filterlim_compose) hoelzl@50327: then have "((\x. norm (1 - g a * inverse (g x))) ---> norm (1 - g a * 0)) (at_right 0)" hoelzl@50327: by (intro tendsto_intros) hoelzl@50327: then have "((\x. norm (1 - g a / g x)) ---> 1) (at_right 0)" hoelzl@50327: by (simp add: inverse_eq_divide) hoelzl@50327: from this[unfolded tendsto_iff, rule_format, of 1] hoelzl@50327: have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" hoelzl@50327: by (auto elim!: eventually_elim1 simp: dist_real_def) hoelzl@50327: hoelzl@50327: moreover hoelzl@50327: from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) ---> norm ((f a - x * g a) * 0)) (at_right 0)" hoelzl@50327: by (intro tendsto_intros) hoelzl@50327: then have "((\t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)" hoelzl@50327: by (simp add: inverse_eq_divide) hoelzl@50327: from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e` hoelzl@50327: have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" hoelzl@50327: by (auto simp: dist_real_def) hoelzl@50327: hoelzl@50327: ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" hoelzl@50327: proof eventually_elim hoelzl@50327: fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" hoelzl@50327: assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" hoelzl@50327: hoelzl@50327: have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" hoelzl@50327: using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ wenzelm@53381: then obtain y where [arith]: "t < y" "y < a" wenzelm@53381: and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" wenzelm@53381: by blast wenzelm@53381: from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" hoelzl@50327: using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps) hoelzl@50327: hoelzl@50327: have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" hoelzl@50327: by (simp add: field_simps) hoelzl@50327: have "norm (f t / g t - x) \ hoelzl@50327: norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" hoelzl@50327: unfolding * by (rule norm_triangle_ineq) hoelzl@50327: also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" hoelzl@50327: by (simp add: abs_mult D_eq dist_real_def) hoelzl@50327: also have "\ < (e / 4) * 2 + e / 2" hoelzl@50327: using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto hoelzl@50327: finally show "dist (f t / g t) x < e" hoelzl@50327: by (simp add: dist_real_def) hoelzl@50327: qed hoelzl@50327: qed hoelzl@50327: hoelzl@50330: lemma lhopital_right_at_top: hoelzl@50330: "LIM x at_right x. (g::real \ real) x :> at_top \ hoelzl@50330: eventually (\x. g' x \ 0) (at_right x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at_right x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at_right x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at_right x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at_right x)" hoelzl@50330: unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift hoelzl@50330: by (rule lhopital_right_0_at_top) hoelzl@50330: hoelzl@50330: lemma lhopital_left_at_top: hoelzl@50330: "LIM x at_left x. (g::real \ real) x :> at_top \ hoelzl@50330: eventually (\x. g' x \ 0) (at_left x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at_left x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at_left x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at_left x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at_left x)" hoelzl@50330: unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror hoelzl@56479: by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) hoelzl@50330: hoelzl@50330: lemma lhopital_at_top: hoelzl@50330: "LIM x at x. (g::real \ real) x :> at_top \ hoelzl@50330: eventually (\x. g' x \ 0) (at x) \ hoelzl@50330: eventually (\x. DERIV f x :> f' x) (at x) \ hoelzl@50330: eventually (\x. DERIV g x :> g' x) (at x) \ hoelzl@50330: ((\ x. (f' x / g' x)) ---> y) (at x) \ hoelzl@50330: ((\ x. f x / g x) ---> y) (at x)" hoelzl@50330: unfolding eventually_at_split filterlim_at_split hoelzl@50330: by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) hoelzl@50330: hoelzl@50347: lemma lhospital_at_top_at_top: hoelzl@50347: fixes f g :: "real \ real" hoelzl@50347: assumes g_0: "LIM x at_top. g x :> at_top" hoelzl@50347: assumes g': "eventually (\x. g' x \ 0) at_top" hoelzl@50347: assumes Df: "eventually (\x. DERIV f x :> f' x) at_top" hoelzl@50347: assumes Dg: "eventually (\x. DERIV g x :> g' x) at_top" hoelzl@50347: assumes lim: "((\ x. (f' x / g' x)) ---> x) at_top" hoelzl@50347: shows "((\ x. f x / g x) ---> x) at_top" hoelzl@50347: unfolding filterlim_at_top_to_right hoelzl@50347: proof (rule lhopital_right_0_at_top) hoelzl@50347: let ?F = "\x. f (inverse x)" hoelzl@50347: let ?G = "\x. g (inverse x)" hoelzl@50347: let ?R = "at_right (0::real)" hoelzl@50347: let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" hoelzl@50347: hoelzl@50347: show "LIM x ?R. ?G x :> at_top" hoelzl@50347: using g_0 unfolding filterlim_at_top_to_right . hoelzl@50347: hoelzl@50347: show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" hoelzl@50347: unfolding eventually_at_right_to_top hoelzl@50347: using Dg eventually_ge_at_top[where c="1::real"] hoelzl@50347: apply eventually_elim hoelzl@50347: apply (rule DERIV_cong) hoelzl@50347: apply (rule DERIV_chain'[where f=inverse]) hoelzl@50347: apply (auto intro!: DERIV_inverse) hoelzl@50347: done hoelzl@50347: hoelzl@50347: show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" hoelzl@50347: unfolding eventually_at_right_to_top hoelzl@50347: using Df eventually_ge_at_top[where c="1::real"] hoelzl@50347: apply eventually_elim hoelzl@50347: apply (rule DERIV_cong) hoelzl@50347: apply (rule DERIV_chain'[where f=inverse]) hoelzl@50347: apply (auto intro!: DERIV_inverse) hoelzl@50347: done hoelzl@50347: hoelzl@50347: show "eventually (\x. ?D g' x \ 0) ?R" hoelzl@50347: unfolding eventually_at_right_to_top hoelzl@50347: using g' eventually_ge_at_top[where c="1::real"] hoelzl@50347: by eventually_elim auto hoelzl@50347: hoelzl@50347: show "((\x. ?D f' x / ?D g' x) ---> x) ?R" hoelzl@50347: unfolding filterlim_at_right_to_top hoelzl@50347: apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) hoelzl@50347: using eventually_ge_at_top[where c="1::real"] hoelzl@56479: by eventually_elim simp hoelzl@50347: qed hoelzl@50347: huffman@21164: end