| author | wenzelm | 
| Mon, 11 Nov 2013 21:16:23 +0100 | |
| changeset 54390 | 0e1566512928 | 
| parent 54230 | b1d955791529 | 
| child 55967 | 5dadc93ff3df | 
| permissions | -rw-r--r-- | 
| 21164 | 1 | (* Title : Deriv.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 4 | Author : Brian Huffman | 
| 21164 | 5 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 6 | GMVT by Benjamin Porter, 2005 | |
| 7 | *) | |
| 8 | ||
| 9 | header{* Differentiation *}
 | |
| 10 | ||
| 11 | theory Deriv | |
| 51526 | 12 | imports Limits | 
| 21164 | 13 | begin | 
| 14 | ||
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 15 | definition | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 16 |   -- {* Frechet derivative: D is derivative of function f at x within s *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 17 |   has_derivative :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow>  bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 18 | (infixl "(has'_derivative)" 12) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 19 | where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 20 | "(f has_derivative f') F \<longleftrightarrow> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 21 | (bounded_linear f' \<and> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 22 | ((\<lambda>y. ((f y - f (Lim F (\<lambda>x. x))) - f' (y - Lim F (\<lambda>x. x))) /\<^sub>R norm (y - Lim F (\<lambda>x. x))) ---> 0) F)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 23 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 24 | lemma FDERIV_eq_rhs: "(f has_derivative f') F \<Longrightarrow> f' = g' \<Longrightarrow> (f has_derivative g') F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 25 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 26 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 27 | ML {*
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 28 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 29 | structure FDERIV_Intros = Named_Thms | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 30 | ( | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 31 |   val name = @{binding FDERIV_intros}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 32 | val description = "introduction rules for FDERIV" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 33 | ) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 34 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 35 | *} | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 36 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 37 | setup {*
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 38 | FDERIV_Intros.setup #> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 39 |   Global_Theory.add_thms_dynamic (@{binding FDERIV_eq_intros},
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 40 |     map_filter (try (fn thm => @{thm FDERIV_eq_rhs} OF [thm])) o FDERIV_Intros.get o Context.proof_of);
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 41 | *} | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 42 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 43 | lemma FDERIV_bounded_linear: "(f has_derivative f') F \<Longrightarrow> bounded_linear f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 44 | by (simp add: has_derivative_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 45 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 46 | lemma FDERIV_ident[FDERIV_intros, simp]: "((\<lambda>x. x) has_derivative (\<lambda>x. x)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 47 | by (simp add: has_derivative_def tendsto_const) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 48 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 49 | lemma FDERIV_const[FDERIV_intros, simp]: "((\<lambda>x. c) has_derivative (\<lambda>x. 0)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 50 | by (simp add: has_derivative_def tendsto_const) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 51 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 52 | lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 53 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 54 | lemma (in bounded_linear) FDERIV: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 55 | "(g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f (g x)) has_derivative (\<lambda>x. f (g' x))) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 56 | using assms unfolding has_derivative_def | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 57 | apply safe | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 58 | apply (erule bounded_linear_compose [OF local.bounded_linear]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 59 | apply (drule local.tendsto) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 60 | apply (simp add: local.scaleR local.diff local.add local.zero) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 61 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 62 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 63 | lemmas FDERIV_scaleR_right [FDERIV_intros] = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 64 | bounded_linear.FDERIV [OF bounded_linear_scaleR_right] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 65 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 66 | lemmas FDERIV_scaleR_left [FDERIV_intros] = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 67 | bounded_linear.FDERIV [OF bounded_linear_scaleR_left] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 68 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 69 | lemmas FDERIV_mult_right [FDERIV_intros] = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 70 | bounded_linear.FDERIV [OF bounded_linear_mult_right] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 71 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 72 | lemmas FDERIV_mult_left [FDERIV_intros] = | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 73 | bounded_linear.FDERIV [OF bounded_linear_mult_left] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 74 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 75 | lemma FDERIV_add[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 76 | assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 77 | shows "((\<lambda>x. f x + g x) has_derivative (\<lambda>x. f' x + g' x)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 78 | unfolding has_derivative_def | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 79 | proof safe | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 80 | let ?x = "Lim F (\<lambda>x. x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 81 | let ?D = "\<lambda>f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 82 | have "((\<lambda>x. ?D f f' x + ?D g g' x) ---> (0 + 0)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 83 | using f g by (intro tendsto_add) (auto simp: has_derivative_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 84 | then show "(?D (\<lambda>x. f x + g x) (\<lambda>x. f' x + g' x) ---> 0) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 85 | by (simp add: field_simps scaleR_add_right scaleR_diff_right) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 86 | qed (blast intro: bounded_linear_add f g FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 87 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 88 | lemma FDERIV_setsum[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 89 | assumes f: "\<And>i. i \<in> I \<Longrightarrow> (f i has_derivative f' i) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 90 | shows "((\<lambda>x. \<Sum>i\<in>I. f i x) has_derivative (\<lambda>x. \<Sum>i\<in>I. f' i x)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 91 | proof cases | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 92 | assume "finite I" from this f show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 93 | by induct (simp_all add: f) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 94 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 95 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 96 | lemma FDERIV_minus[simp, FDERIV_intros]: "(f has_derivative f') F \<Longrightarrow> ((\<lambda>x. - f x) has_derivative (\<lambda>x. - f' x)) F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 97 | using FDERIV_scaleR_right[of f f' F "-1"] by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 98 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 99 | lemma FDERIV_diff[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 100 | "(f has_derivative f') F \<Longrightarrow> (g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f x - g x) has_derivative (\<lambda>x. f' x - g' x)) F" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 101 | by (simp only: diff_conv_add_uminus FDERIV_add FDERIV_minus) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 102 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 103 | abbreviation | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 104 |   -- {* Frechet derivative: D is derivative of function f at x within s *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 105 |   FDERIV :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow>  ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 106 |   ("(FDERIV (_)/ (_)/ : (_)/ :> (_))" [1000, 1000, 1000, 60] 60)
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 107 | where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 108 | "FDERIV f x : s :> f' \<equiv> (f has_derivative f') (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 109 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 110 | abbreviation | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 111 | fderiv_at :: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 112 |     "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 113 |     ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 114 | where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 115 | "FDERIV f x :> D \<equiv> FDERIV f x : UNIV :> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 116 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 117 | lemma FDERIV_def: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 118 | "FDERIV f x : s :> f' \<longleftrightarrow> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 119 | (bounded_linear f' \<and> ((\<lambda>y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) ---> 0) (at x within s))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 120 | by (cases "at x within s = bot") (simp_all add: has_derivative_def Lim_ident_at) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 121 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 122 | lemma FDERIV_iff_norm: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 123 | "FDERIV f x : s :> f' \<longleftrightarrow> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 124 | (bounded_linear f' \<and> ((\<lambda>y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) ---> 0) (at x within s))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 125 | using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 126 | by (simp add: FDERIV_def divide_inverse ac_simps) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 127 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 128 | lemma fderiv_def: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 129 | "FDERIV f x :> D = (bounded_linear D \<and> (\<lambda>h. norm (f (x + h) - f x - D h) / norm h) -- 0 --> 0)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 130 | unfolding FDERIV_iff_norm LIM_offset_zero_iff[of _ _ x] by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 131 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 132 | lemma field_fderiv_def: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 133 | fixes x :: "'a::real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 134 | shows "FDERIV f x :> (\<lambda>h. h * D) = (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 135 | apply (unfold fderiv_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 136 | apply (simp add: bounded_linear_mult_left) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 137 | apply (simp cong: LIM_cong add: nonzero_norm_divide [symmetric]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 138 | apply (subst diff_divide_distrib) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 139 | apply (subst times_divide_eq_left [symmetric]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 140 | apply (simp cong: LIM_cong) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 141 | apply (simp add: tendsto_norm_zero_iff LIM_zero_iff) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 142 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 143 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 144 | lemma FDERIV_I: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 145 | "bounded_linear f' \<Longrightarrow> ((\<lambda>y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) ---> 0) (at x within s) \<Longrightarrow> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 146 | FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 147 | by (simp add: FDERIV_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 148 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 149 | lemma FDERIV_I_sandwich: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 150 | assumes e: "0 < e" and bounded: "bounded_linear f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 151 | and sandwich: "(\<And>y. y \<in> s \<Longrightarrow> y \<noteq> x \<Longrightarrow> dist y x < e \<Longrightarrow> norm ((f y - f x) - f' (y - x)) / norm (y - x) \<le> H y)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 152 | and "(H ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 153 | shows "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 154 | unfolding FDERIV_iff_norm | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 155 | proof safe | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 156 | show "((\<lambda>y. norm (f y - f x - f' (y - x)) / norm (y - x)) ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 157 | proof (rule tendsto_sandwich[where f="\<lambda>x. 0"]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 158 | show "(H ---> 0) (at x within s)" by fact | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 159 | show "eventually (\<lambda>n. norm (f n - f x - f' (n - x)) / norm (n - x) \<le> H n) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 160 | unfolding eventually_at using e sandwich by auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 161 | qed (auto simp: le_divide_eq tendsto_const) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 162 | qed fact | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 163 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 164 | lemma FDERIV_subset: "FDERIV f x : s :> f' \<Longrightarrow> t \<subseteq> s \<Longrightarrow> FDERIV f x : t :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 165 | by (auto simp add: FDERIV_iff_norm intro: tendsto_within_subset) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 166 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 167 | subsection {* Continuity *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 168 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 169 | lemma FDERIV_continuous: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 170 | assumes f: "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 171 | shows "continuous (at x within s) f" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 172 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 173 | from f interpret F: bounded_linear f' by (rule FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 174 | note F.tendsto[tendsto_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 175 | let ?L = "\<lambda>f. (f ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 176 | have "?L (\<lambda>y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 177 | using f unfolding FDERIV_iff_norm by blast | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 178 | then have "?L (\<lambda>y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 179 | by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 180 | also have "?m \<longleftrightarrow> ?L (\<lambda>y. norm ((f y - f x) - f' (y - x)))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 181 | by (intro filterlim_cong) (simp_all add: eventually_at_filter) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 182 | finally have "?L (\<lambda>y. (f y - f x) - f' (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 183 | by (rule tendsto_norm_zero_cancel) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 184 | then have "?L (\<lambda>y. ((f y - f x) - f' (y - x)) + f' (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 185 | by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 186 | then have "?L (\<lambda>y. f y - f x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 187 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 188 | from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 189 | by (simp add: continuous_within) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 190 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 191 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 192 | subsection {* Composition *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 193 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 194 | lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \<Longrightarrow> (f ---> y) (at x within s) \<longleftrightarrow> (f ---> y) (inf (nhds x) (principal s))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 195 | unfolding tendsto_def eventually_inf_principal eventually_at_filter | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 196 | by (intro ext all_cong imp_cong) (auto elim!: eventually_elim1) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 197 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 198 | lemma FDERIV_in_compose: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 199 | assumes f: "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 200 | assumes g: "FDERIV g (f x) : (f`s) :> g'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 201 | shows "FDERIV (\<lambda>x. g (f x)) x : s :> (\<lambda>x. g' (f' x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 202 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 203 | from f interpret F: bounded_linear f' by (rule FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 204 | from g interpret G: bounded_linear g' by (rule FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 205 | from F.bounded obtain kF where kF: "\<And>x. norm (f' x) \<le> norm x * kF" by fast | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 206 | from G.bounded obtain kG where kG: "\<And>x. norm (g' x) \<le> norm x * kG" by fast | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 207 | note G.tendsto[tendsto_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 208 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 209 | let ?L = "\<lambda>f. (f ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 210 | let ?D = "\<lambda>f f' x y. (f y - f x) - f' (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 211 | let ?N = "\<lambda>f f' x y. norm (?D f f' x y) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 212 | let ?gf = "\<lambda>x. g (f x)" and ?gf' = "\<lambda>x. g' (f' x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 213 | def Nf \<equiv> "?N f f' x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 214 | def Ng \<equiv> "\<lambda>y. ?N g g' (f x) (f y)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 215 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 216 | show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 217 | proof (rule FDERIV_I_sandwich[of 1]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 218 | show "bounded_linear (\<lambda>x. g' (f' x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 219 | using f g by (blast intro: bounded_linear_compose FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 220 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 221 | fix y::'a assume neq: "y \<noteq> x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 222 | have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 223 | by (simp add: G.diff G.add field_simps) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 224 | also have "\<dots> \<le> norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 225 | by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 226 | also have "\<dots> \<le> Nf y * kG + Ng y * (Nf y + kF)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 227 | proof (intro add_mono mult_left_mono) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 228 | have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 229 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 230 | also have "\<dots> \<le> norm (?D f f' x y) + norm (f' (y - x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 231 | by (rule norm_triangle_ineq) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 232 | also have "\<dots> \<le> norm (?D f f' x y) + norm (y - x) * kF" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 233 | using kF by (intro add_mono) simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 234 | finally show "norm (f y - f x) / norm (y - x) \<le> Nf y + kF" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 235 | by (simp add: neq Nf_def field_simps) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 236 | qed (insert kG, simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 237 | finally show "?N ?gf ?gf' x y \<le> Nf y * kG + Ng y * (Nf y + kF)" . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 238 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 239 | have [tendsto_intros]: "?L Nf" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 240 | using f unfolding FDERIV_iff_norm Nf_def .. | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 241 | from f have "(f ---> f x) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 242 | by (blast intro: FDERIV_continuous continuous_within[THEN iffD1]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 243 | then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 244 | unfolding filterlim_def | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 245 | by (simp add: eventually_filtermap eventually_at_filter le_principal) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 246 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 247 | have "((?N g g' (f x)) ---> 0) (at (f x) within f`s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 248 | using g unfolding FDERIV_iff_norm .. | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 249 | then have g': "((?N g g' (f x)) ---> 0) (inf (nhds (f x)) (principal (f`s)))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 250 | by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 251 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 252 | have [tendsto_intros]: "?L Ng" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 253 | unfolding Ng_def by (rule filterlim_compose[OF g' f']) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 254 | show "((\<lambda>y. Nf y * kG + Ng y * (Nf y + kF)) ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 255 | by (intro tendsto_eq_intros) auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 256 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 257 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 258 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 259 | lemma FDERIV_compose: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 260 | "FDERIV f x : s :> f' \<Longrightarrow> FDERIV g (f x) :> g' \<Longrightarrow> FDERIV (\<lambda>x. g (f x)) x : s :> (\<lambda>x. g' (f' x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 261 | by (blast intro: FDERIV_in_compose FDERIV_subset) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 262 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 263 | lemma (in bounded_bilinear) FDERIV: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 264 | assumes f: "FDERIV f x : s :> f'" and g: "FDERIV g x : s :> g'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 265 | shows "FDERIV (\<lambda>x. f x ** g x) x : s :> (\<lambda>h. f x ** g' h + f' h ** g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 266 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 267 | from bounded_linear.bounded [OF FDERIV_bounded_linear [OF f]] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 268 | obtain KF where norm_F: "\<And>x. norm (f' x) \<le> norm x * KF" by fast | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 269 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 270 | from pos_bounded obtain K where K: "0 < K" and norm_prod: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 271 | "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" by fast | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 272 | let ?D = "\<lambda>f f' y. f y - f x - f' (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 273 | let ?N = "\<lambda>f f' y. norm (?D f f' y) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 274 | def Ng =="?N g g'" and Nf =="?N f f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 275 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 276 | let ?fun1 = "\<lambda>y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 277 | let ?fun2 = "\<lambda>y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 278 | let ?F = "at x within s" | 
| 21164 | 279 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 280 | show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 281 | proof (rule FDERIV_I_sandwich[of 1]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 282 | show "bounded_linear (\<lambda>h. f x ** g' h + f' h ** g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 283 | by (intro bounded_linear_add | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 284 | bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 285 | FDERIV_bounded_linear [OF g] FDERIV_bounded_linear [OF f]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 286 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 287 | from g have "(g ---> g x) ?F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 288 | by (intro continuous_within[THEN iffD1] FDERIV_continuous) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 289 | moreover from f g have "(Nf ---> 0) ?F" "(Ng ---> 0) ?F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 290 | by (simp_all add: FDERIV_iff_norm Ng_def Nf_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 291 | ultimately have "(?fun2 ---> norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 292 | by (intro tendsto_intros) (simp_all add: LIM_zero_iff) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 293 | then show "(?fun2 ---> 0) ?F" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 294 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 295 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 296 | fix y::'d assume "y \<noteq> x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 297 | 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)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 298 | by (simp add: diff_left diff_right add_left add_right field_simps) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 299 | also have "\<dots> \<le> (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 300 | norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 301 | by (intro divide_right_mono mult_mono' | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 302 | order_trans [OF norm_triangle_ineq add_mono] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 303 | order_trans [OF norm_prod mult_right_mono] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 304 | mult_nonneg_nonneg order_refl norm_ge_zero norm_F | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 305 | K [THEN order_less_imp_le]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 306 | also have "\<dots> = ?fun2 y" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 307 | by (simp add: add_divide_distrib Ng_def Nf_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 308 | finally show "?fun1 y \<le> ?fun2 y" . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 309 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 310 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 311 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 312 | lemmas FDERIV_mult[simp, FDERIV_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 313 | lemmas FDERIV_scaleR[simp, FDERIV_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 314 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 315 | lemma FDERIV_setprod[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 316 | fixes f :: "'i \<Rightarrow> 'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 317 | assumes f: "\<And>i. i \<in> I \<Longrightarrow> FDERIV (f i) x : s :> f' i" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 318 |   shows "FDERIV (\<lambda>x. \<Prod>i\<in>I. f i x) x : s :> (\<lambda>y. \<Sum>i\<in>I. f' i y * (\<Prod>j\<in>I - {i}. f j x))"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 319 | proof cases | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 320 | assume "finite I" from this f show ?thesis | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 321 | proof induct | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 322 | case (insert i I) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 323 |     let ?P = "\<lambda>y. f i x * (\<Sum>i\<in>I. f' i y * (\<Prod>j\<in>I - {i}. f j x)) + (f' i y) * (\<Prod>i\<in>I. f i x)"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 324 | have "FDERIV (\<lambda>x. f i x * (\<Prod>i\<in>I. f i x)) x : s :> ?P" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 325 | using insert by (intro FDERIV_mult) auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 326 |     also have "?P = (\<lambda>y. \<Sum>i'\<in>insert i I. f' i' y * (\<Prod>j\<in>insert i I - {i'}. f j x))"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 327 | using insert(1,2) by (auto simp add: setsum_right_distrib insert_Diff_if intro!: ext setsum_cong) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 328 | finally show ?case | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 329 | using insert by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 330 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 331 | qed simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 332 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 333 | lemma FDERIV_power[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 334 | fixes f :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 335 | assumes f: "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 336 | shows "FDERIV (\<lambda>x. f x^n) x : s :> (\<lambda>y. of_nat n * f' y * f x^(n - 1))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 337 |   using FDERIV_setprod[OF f, of "{..< n}"] by (simp add: setprod_constant ac_simps)
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 338 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 339 | lemma FDERIV_inverse': | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 340 | fixes x :: "'a::real_normed_div_algebra" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 341 | assumes x: "x \<noteq> 0" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 342 | shows "FDERIV inverse x : s :> (\<lambda>h. - (inverse x * h * inverse x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 343 | (is "FDERIV ?inv x : s :> ?f") | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 344 | proof (rule FDERIV_I_sandwich) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 345 | show "bounded_linear (\<lambda>h. - (?inv x * h * ?inv x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 346 | apply (rule bounded_linear_minus) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 347 | apply (rule bounded_linear_mult_const) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 348 | apply (rule bounded_linear_const_mult) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 349 | apply (rule bounded_linear_ident) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 350 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 351 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 352 | show "0 < norm x" using x by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 353 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 354 | show "((\<lambda>y. norm (?inv y - ?inv x) * norm (?inv x)) ---> 0) (at x within s)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 355 | apply (rule tendsto_mult_left_zero) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 356 | apply (rule tendsto_norm_zero) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 357 | apply (rule LIM_zero) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 358 | apply (rule tendsto_inverse) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 359 | apply (rule tendsto_ident_at) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 360 | apply (rule x) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 361 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 362 | next | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 363 | fix y::'a assume h: "y \<noteq> x" "dist y x < norm x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 364 | then have "y \<noteq> 0" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 365 | by (auto simp: norm_conv_dist dist_commute) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 366 | have "norm (?inv y - ?inv x - ?f (y -x)) / norm (y - x) = norm ((?inv y - ?inv x) * (y - x) * ?inv x) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 367 | apply (subst inverse_diff_inverse [OF `y \<noteq> 0` x]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 368 | apply (subst minus_diff_minus) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 369 | apply (subst norm_minus_cancel) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 370 | apply (simp add: left_diff_distrib) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 371 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 372 | also have "\<dots> \<le> norm (?inv y - ?inv x) * norm (y - x) * norm (?inv x) / norm (y - x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 373 | apply (rule divide_right_mono [OF _ norm_ge_zero]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 374 | apply (rule order_trans [OF norm_mult_ineq]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 375 | apply (rule mult_right_mono [OF _ norm_ge_zero]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 376 | apply (rule norm_mult_ineq) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 377 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 378 | also have "\<dots> = norm (?inv y - ?inv x) * norm (?inv x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 379 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 380 | finally show "norm (?inv y - ?inv x - ?f (y -x)) / norm (y - x) \<le> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 381 | norm (?inv y - ?inv x) * norm (?inv x)" . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 382 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 383 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 384 | lemma FDERIV_inverse[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 385 | fixes f :: "_ \<Rightarrow> 'a::real_normed_div_algebra" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 386 | assumes x: "f x \<noteq> 0" and f: "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 387 | shows "FDERIV (\<lambda>x. inverse (f x)) x : s :> (\<lambda>h. - (inverse (f x) * f' h * inverse (f x)))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 388 | using FDERIV_compose[OF f FDERIV_inverse', OF x] . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 389 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 390 | lemma FDERIV_divide[simp, FDERIV_intros]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 391 | fixes f :: "_ \<Rightarrow> 'a::real_normed_div_algebra" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 392 | assumes g: "FDERIV g x : s :> g'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 393 | assumes x: "f x \<noteq> 0" and f: "FDERIV f x : s :> f'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 394 | shows "FDERIV (\<lambda>x. g x / f x) x : s :> (\<lambda>h. - g x * (inverse (f x) * f' h * inverse (f x)) + g' h / f x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 395 | using FDERIV_mult[OF g FDERIV_inverse[OF x f]] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 396 | by (simp add: divide_inverse) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 397 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 398 | subsection {* Uniqueness *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 399 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 400 | text {*
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 401 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 402 | This can not generally shown for @{const FDERIV}, as we need to approach the point from
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 403 | all directions. There is a proof in @{text Multivariate_Analysis} for @{text euclidean_space}.
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 404 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 405 | *} | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 406 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 407 | lemma FDERIV_zero_unique: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 408 | assumes "FDERIV (\<lambda>x. 0) x :> F" shows "F = (\<lambda>h. 0)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 409 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 410 | interpret F: bounded_linear F | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 411 | using assms by (rule FDERIV_bounded_linear) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 412 | let ?r = "\<lambda>h. norm (F h) / norm h" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 413 | have *: "?r -- 0 --> 0" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 414 | using assms unfolding fderiv_def by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 415 | show "F = (\<lambda>h. 0)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 416 | proof | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 417 | fix h show "F h = 0" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 418 | proof (rule ccontr) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 419 | assume **: "F h \<noteq> 0" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 420 | then have h: "h \<noteq> 0" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 421 | by (clarsimp simp add: F.zero) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 422 | with ** have "0 < ?r h" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 423 | by (simp add: divide_pos_pos) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 424 | from LIM_D [OF * this] obtain s where s: "0 < s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 425 | and r: "\<And>x. x \<noteq> 0 \<Longrightarrow> norm x < s \<Longrightarrow> ?r x < ?r h" by auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 426 | from dense [OF s] obtain t where t: "0 < t \<and> t < s" .. | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 427 | let ?x = "scaleR (t / norm h) h" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 428 | have "?x \<noteq> 0" and "norm ?x < s" using t h by simp_all | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 429 | hence "?r ?x < ?r h" by (rule r) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 430 | thus "False" using t h by (simp add: F.scaleR) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 431 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 432 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 433 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 434 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 435 | lemma FDERIV_unique: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 436 | assumes "FDERIV f x :> F" and "FDERIV f x :> F'" shows "F = F'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 437 | proof - | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 438 | have "FDERIV (\<lambda>x. 0) x :> (\<lambda>h. F h - F' h)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 439 | using FDERIV_diff [OF assms] by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 440 | hence "(\<lambda>h. F h - F' h) = (\<lambda>h. 0)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 441 | by (rule FDERIV_zero_unique) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 442 | thus "F = F'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 443 | unfolding fun_eq_iff right_minus_eq . | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 444 | qed | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 445 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 446 | subsection {* Differentiability predicate *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 447 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 448 | definition isDiff :: "'a filter \<Rightarrow> ('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool" where
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 449 | isDiff_def: "isDiff F f \<longleftrightarrow> (\<exists>D. (f has_derivative D) F)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 450 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 451 | abbreviation differentiable_in :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 452 |     ("(_) differentiable (_) in (_)"  [1000, 1000, 60] 60) where
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 453 | "f differentiable x in s \<equiv> isDiff (at x within s) f" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 454 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 455 | abbreviation differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 456 | (infixl "differentiable" 60) where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 457 | "f differentiable x \<equiv> f differentiable x in UNIV" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 458 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 459 | lemma differentiable_subset: "f differentiable x in s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f differentiable x in t" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 460 | unfolding isDiff_def by (blast intro: FDERIV_subset) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 461 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 462 | lemma differentiable_ident [simp]: "isDiff F (\<lambda>x. x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 463 | unfolding isDiff_def by (blast intro: FDERIV_ident) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 464 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 465 | lemma differentiable_const [simp]: "isDiff F (\<lambda>z. a)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 466 | unfolding isDiff_def by (blast intro: FDERIV_const) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 467 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 468 | lemma differentiable_in_compose: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 469 | "f differentiable (g x) in (g`s) \<Longrightarrow> g differentiable x in s \<Longrightarrow> (\<lambda>x. f (g x)) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 470 | unfolding isDiff_def by (blast intro: FDERIV_in_compose) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 471 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 472 | lemma differentiable_compose: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 473 | "f differentiable (g x) \<Longrightarrow> g differentiable x in s \<Longrightarrow> (\<lambda>x. f (g x)) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 474 | by (blast intro: differentiable_in_compose differentiable_subset) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 475 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 476 | lemma differentiable_sum [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 477 | "isDiff F f \<Longrightarrow> isDiff F g \<Longrightarrow> isDiff F (\<lambda>x. f x + g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 478 | unfolding isDiff_def by (blast intro: FDERIV_add) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 479 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 480 | lemma differentiable_minus [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 481 | "isDiff F f \<Longrightarrow> isDiff F (\<lambda>x. - f x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 482 | unfolding isDiff_def by (blast intro: FDERIV_minus) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 483 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 484 | lemma differentiable_diff [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 485 | "isDiff F f \<Longrightarrow> isDiff F g \<Longrightarrow> isDiff F (\<lambda>x. f x - g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 486 | unfolding isDiff_def by (blast intro: FDERIV_diff) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 487 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 488 | lemma differentiable_mult [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 489 | fixes f g :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_algebra" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 490 | shows "f differentiable x in s \<Longrightarrow> g differentiable x in s \<Longrightarrow> (\<lambda>x. f x * g x) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 491 | unfolding isDiff_def by (blast intro: FDERIV_mult) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 492 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 493 | lemma differentiable_inverse [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 494 | fixes f :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 495 | shows "f differentiable x in s \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> (\<lambda>x. inverse (f x)) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 496 | unfolding isDiff_def by (blast intro: FDERIV_inverse) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 497 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 498 | lemma differentiable_divide [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 499 | fixes f g :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 500 | shows "f differentiable x in s \<Longrightarrow> g differentiable x in s \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> (\<lambda>x. f x / g x) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 501 | unfolding divide_inverse using assms by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 502 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 503 | lemma differentiable_power [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 504 | fixes f g :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 505 | shows "f differentiable x in s \<Longrightarrow> (\<lambda>x. f x ^ n) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 506 | unfolding isDiff_def by (blast intro: FDERIV_power) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 507 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 508 | lemma differentiable_scaleR [simp]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 509 | "f differentiable x in s \<Longrightarrow> g differentiable x in s \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 510 | unfolding isDiff_def by (blast intro: FDERIV_scaleR) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 511 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 512 | definition | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 513 |   -- {*Differentiation: D is derivative of function f at x*}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 514 | deriv :: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 515 |     "('a::real_normed_field \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 516 |     ("(DERIV (_)/ (_)/ : (_)/ :> (_))" [1000, 1000, 1000, 60] 60)
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 517 | where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 518 | deriv_fderiv: "DERIV f x : s :> D = FDERIV f x : s :> (\<lambda>x. x * D)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 519 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 520 | abbreviation | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 521 |   -- {*Differentiation: D is derivative of function f at x*}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 522 | deriv_at :: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 523 |     "('a::real_normed_field \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 524 |     ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 525 | where | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 526 | "DERIV f x :> D \<equiv> DERIV f x : UNIV :> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 527 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 528 | lemma differentiable_def: "(f::real \<Rightarrow> real) differentiable x in s \<longleftrightarrow> (\<exists>D. DERIV f x : s :> D)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 529 | proof safe | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 530 | assume "f differentiable x in s" | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 531 | then obtain f' where *: "FDERIV f x : s :> f'" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 532 | unfolding isDiff_def by auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 533 | then obtain c where "f' = (\<lambda>x. x * c)" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 534 | by (metis real_bounded_linear FDERIV_bounded_linear) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 535 | with * show "\<exists>D. DERIV f x : s :> D" | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 536 | unfolding deriv_fderiv by auto | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 537 | qed (auto simp: isDiff_def deriv_fderiv) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 538 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 539 | lemma differentiableE [elim?]: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 540 | fixes f :: "real \<Rightarrow> real" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 541 | assumes f: "f differentiable x in s" obtains df where "DERIV f x : s :> df" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 542 | using assms by (auto simp: differentiable_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 543 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 544 | lemma differentiableD: "(f::real \<Rightarrow> real) differentiable x in s \<Longrightarrow> \<exists>D. DERIV f x : s :> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 545 | by (auto elim: differentiableE) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 546 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 547 | lemma differentiableI: "DERIV f x : s :> D \<Longrightarrow> (f::real \<Rightarrow> real) differentiable x in s" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 548 | by (force simp add: differentiable_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 549 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 550 | lemma DERIV_I_FDERIV: "FDERIV f x : s :> F \<Longrightarrow> (\<And>x. x * F' = F x) \<Longrightarrow> DERIV f x : s :> F'" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 551 | by (simp add: deriv_fderiv) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 552 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 553 | lemma DERIV_D_FDERIV: "DERIV f x : s :> F \<Longrightarrow> FDERIV f x : s :> (\<lambda>x. x * F)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 554 | by (simp add: deriv_fderiv) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 555 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 556 | lemma deriv_def: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 557 | "DERIV f x :> D \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 558 | apply (simp add: deriv_fderiv fderiv_def bounded_linear_mult_left LIM_zero_iff[symmetric, of _ D]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 559 | apply (subst (2) tendsto_norm_zero_iff[symmetric]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 560 | apply (rule filterlim_cong) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 561 | apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 562 | done | 
| 21164 | 563 | |
| 564 | subsection {* Derivatives *}
 | |
| 565 | ||
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 566 | lemma DERIV_iff: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 567 | by (simp add: deriv_def) | 
| 21164 | 568 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 569 | lemma DERIV_D: "DERIV f x :> D \<Longrightarrow> (\<lambda>h. (f (x + h) - f x) / h) -- 0 --> D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 570 | by (simp add: deriv_def) | 
| 21164 | 571 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 572 | lemma DERIV_const [simp]: "DERIV (\<lambda>x. k) x : s :> 0" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 573 | by (rule DERIV_I_FDERIV[OF FDERIV_const]) auto | 
| 21164 | 574 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 575 | lemma DERIV_ident [simp]: "DERIV (\<lambda>x. x) x : s :> 1" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 576 | by (rule DERIV_I_FDERIV[OF FDERIV_ident]) auto | 
| 21164 | 577 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 578 | lemma DERIV_add: "DERIV f x : s :> D \<Longrightarrow> DERIV g x : s :> E \<Longrightarrow> DERIV (\<lambda>x. f x + g x) x : s :> D + E" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 579 | by (rule DERIV_I_FDERIV[OF FDERIV_add]) (auto simp: field_simps dest: DERIV_D_FDERIV) | 
| 21164 | 580 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 581 | lemma DERIV_minus: "DERIV f x : s :> D \<Longrightarrow> DERIV (\<lambda>x. - f x) x : s :> - D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 582 | by (rule DERIV_I_FDERIV[OF FDERIV_minus]) (auto simp: field_simps dest: DERIV_D_FDERIV) | 
| 21164 | 583 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 584 | lemma DERIV_diff: "DERIV f x : s :> D \<Longrightarrow> DERIV g x : s :> E \<Longrightarrow> DERIV (\<lambda>x. f x - g x) x : s :> D - E" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 585 | by (rule DERIV_I_FDERIV[OF FDERIV_diff]) (auto simp: field_simps dest: DERIV_D_FDERIV) | 
| 21164 | 586 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 587 | lemma DERIV_add_minus: "DERIV f x : s :> D \<Longrightarrow> DERIV g x : s :> E \<Longrightarrow> DERIV (\<lambda>x. f x + - g x) x : s :> D + - E" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 588 | by (simp only: DERIV_add DERIV_minus) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 589 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 590 | lemma DERIV_continuous: "DERIV f x : s :> D \<Longrightarrow> continuous (at x within s) f" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 591 | by (drule FDERIV_continuous[OF DERIV_D_FDERIV]) simp | 
| 21164 | 592 | |
| 593 | lemma DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x" | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 594 | by (auto dest!: DERIV_continuous) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 595 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 596 | lemma DERIV_mult': "DERIV f x : s :> D \<Longrightarrow> DERIV g x : s :> E \<Longrightarrow> DERIV (\<lambda>x. f x * g x) x : s :> f x * E + D * g x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 597 | by (rule DERIV_I_FDERIV[OF FDERIV_mult]) (auto simp: field_simps dest: DERIV_D_FDERIV) | 
| 21164 | 598 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 599 | lemma DERIV_mult: "DERIV f x : s :> Da \<Longrightarrow> DERIV g x : s :> Db \<Longrightarrow> DERIV (\<lambda>x. f x * g x) x : s :> Da * g x + Db * f x" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 600 | by (rule DERIV_I_FDERIV[OF FDERIV_mult]) (auto simp: field_simps dest: DERIV_D_FDERIV) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 601 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 602 | text {* Derivative of linear multiplication *}
 | 
| 21164 | 603 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 604 | lemma DERIV_cmult: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 605 | "DERIV f x : s :> D ==> DERIV (%x. c * f x) x : s :> c*D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 606 | by (drule DERIV_mult' [OF DERIV_const], simp) | 
| 21164 | 607 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 608 | lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x : s :> c" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 609 | by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 610 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 611 | lemma DERIV_cdivide: "DERIV f x : s :> D ==> DERIV (%x. f x / c) x : s :> D / c" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 612 | apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x : s :> (1 / c) * D", force) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 613 | apply (erule DERIV_cmult) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 614 | done | 
| 21164 | 615 | |
| 616 | lemma DERIV_unique: | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 617 | "DERIV f x :> D \<Longrightarrow> DERIV f x :> E \<Longrightarrow> D = E" | 
| 50331 | 618 | unfolding deriv_def by (rule LIM_unique) | 
| 21164 | 619 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 620 | lemma DERIV_setsum': | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 621 | "(\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x : s :> (f' x n)) \<Longrightarrow> DERIV (\<lambda>x. setsum (f x) S) x : s :> setsum (f' x) S" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 622 | by (rule DERIV_I_FDERIV[OF FDERIV_setsum]) (auto simp: setsum_right_distrib dest: DERIV_D_FDERIV) | 
| 21164 | 623 | |
| 31880 | 624 | lemma DERIV_setsum: | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 625 | "finite S \<Longrightarrow> (\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x : s :> (f' x n)) \<Longrightarrow> DERIV (\<lambda>x. setsum (f x) S) x : s :> setsum (f' x) S" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 626 | by (rule DERIV_setsum') | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 627 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 628 | lemma DERIV_sumr [rule_format (no_asm)]: (* REMOVE *) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 629 | "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x : s :> (f' r x)) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 630 | --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x : s :> (\<Sum>r=m..<n. f' r x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 631 | by (auto intro: DERIV_setsum) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 632 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 633 | lemma DERIV_inverse': | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 634 | "DERIV f x : s :> D \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse (f x)) x : s :> - (inverse (f x) * D * inverse (f x))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 635 | by (rule DERIV_I_FDERIV[OF FDERIV_inverse]) (auto dest: DERIV_D_FDERIV) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 636 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 637 | text {* Power of @{text "-1"} *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 638 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 639 | lemma DERIV_inverse: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 640 | "x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse(x)) x : s :> - (inverse x ^ Suc (Suc 0))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 641 | by (drule DERIV_inverse' [OF DERIV_ident]) simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 642 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 643 | text {* Derivative of inverse *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 644 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 645 | lemma DERIV_inverse_fun: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 646 | "DERIV f x : s :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse (f x)) x : s :> (- (d * inverse(f x ^ Suc (Suc 0))))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 647 | by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 648 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 649 | text {* Derivative of quotient *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 650 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 651 | lemma DERIV_divide: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 652 | "DERIV f x : s :> D \<Longrightarrow> DERIV g x : s :> E \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. f x / g x) x : s :> (D * g x - f x * E) / (g x * g x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 653 | by (rule DERIV_I_FDERIV[OF FDERIV_divide]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 654 | (auto dest: DERIV_D_FDERIV simp: field_simps nonzero_inverse_mult_distrib divide_inverse) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 655 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 656 | lemma DERIV_quotient: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 657 | "DERIV f x : s :> d \<Longrightarrow> DERIV g x : s :> e \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>y. f y / g y) x : s :> (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 658 | by (drule (2) DERIV_divide) (simp add: mult_commute) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 659 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 660 | lemma DERIV_power_Suc: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 661 | "DERIV f x : s :> D \<Longrightarrow> DERIV (\<lambda>x. f x ^ Suc n) x : s :> (1 + of_nat n) * (D * f x ^ n)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 662 | by (rule DERIV_I_FDERIV[OF FDERIV_power]) (auto simp: deriv_fderiv) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 663 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 664 | lemma DERIV_power: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 665 | "DERIV f x : s :> D \<Longrightarrow> DERIV (\<lambda>x. f x ^ n) x : s :> of_nat n * (D * f x ^ (n - Suc 0))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 666 | by (rule DERIV_I_FDERIV[OF FDERIV_power]) (auto simp: deriv_fderiv) | 
| 31880 | 667 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 668 | lemma DERIV_pow: "DERIV (%x. x ^ n) x : s :> real n * (x ^ (n - Suc 0))" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 669 | apply (cut_tac DERIV_power [OF DERIV_ident]) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 670 | apply (simp add: real_of_nat_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 671 | done | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 672 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 673 | lemma DERIV_chain': "DERIV f x : s :> D \<Longrightarrow> DERIV g (f x) :> E \<Longrightarrow> DERIV (\<lambda>x. g (f x)) x : s :> E * D" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 674 | using FDERIV_compose[of f "\<lambda>x. x * D" x s g "\<lambda>x. x * E"] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 675 | by (auto simp: deriv_fderiv ac_simps dest: FDERIV_subset) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 676 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 677 | text {* Standard version *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 678 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 679 | lemma DERIV_chain: "DERIV f (g x) :> Da \<Longrightarrow> DERIV g x : s :> Db \<Longrightarrow> DERIV (f o g) x : s :> Da * Db" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 680 | by (drule (1) DERIV_chain', simp add: o_def mult_commute) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 681 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 682 | lemma DERIV_chain2: "DERIV f (g x) :> Da \<Longrightarrow> DERIV g x : s :> Db \<Longrightarrow> DERIV (%x. f (g x)) x : s :> Da * Db" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 683 | by (auto dest: DERIV_chain simp add: o_def) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 684 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 685 | subsubsection {* @{text "DERIV_intros"} *}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 686 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 687 | ML {*
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 688 | structure Deriv_Intros = Named_Thms | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 689 | ( | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 690 |   val name = @{binding DERIV_intros}
 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 691 | val description = "DERIV introduction rules" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 692 | ) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 693 | *} | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 694 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 695 | setup Deriv_Intros.setup | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 696 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 697 | lemma DERIV_cong: "DERIV f x : s :> X \<Longrightarrow> X = Y \<Longrightarrow> DERIV f x : s :> Y" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 698 | by simp | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 699 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 700 | declare | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 701 | DERIV_const[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 702 | DERIV_ident[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 703 | DERIV_add[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 704 | DERIV_minus[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 705 | DERIV_mult[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 706 | DERIV_diff[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 707 | DERIV_inverse'[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 708 | DERIV_divide[THEN DERIV_cong, DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 709 | DERIV_power[where 'a=real, THEN DERIV_cong, | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 710 | unfolded real_of_nat_def[symmetric], DERIV_intros] | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 711 | DERIV_setsum[THEN DERIV_cong, DERIV_intros] | 
| 21164 | 712 | |
| 713 | text{*Alternative definition for differentiability*}
 | |
| 714 | ||
| 715 | lemma DERIV_LIM_iff: | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 716 |   fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 717 | "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) = | 
| 21164 | 718 | ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" | 
| 719 | apply (rule iffI) | |
| 720 | apply (drule_tac k="- a" in LIM_offset) | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 721 | apply simp | 
| 21164 | 722 | apply (drule_tac k="a" in LIM_offset) | 
| 723 | apply (simp add: add_commute) | |
| 724 | done | |
| 725 | ||
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 726 | lemma DERIV_iff2: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) --x --> D" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 727 | by (simp add: deriv_def DERIV_LIM_iff) | 
| 21164 | 728 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 729 | lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow> | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 730 | DERIV f x :> u \<longleftrightarrow> DERIV g y :> v" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 731 | unfolding DERIV_iff2 | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 732 | proof (rule filterlim_cong) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 733 | assume *: "eventually (\<lambda>x. f x = g x) (nhds x)" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 734 | moreover from * have "f x = g x" by (auto simp: eventually_nhds) | 
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 735 | moreover assume "x = y" "u = v" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 736 | ultimately show "eventually (\<lambda>xa. (f xa - f x) / (xa - x) = (g xa - g y) / (xa - y)) (at x)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 737 | by (auto simp: eventually_at_filter elim: eventually_elim1) | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 738 | qed simp_all | 
| 21164 | 739 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 740 | lemma DERIV_shift: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 741 | "(DERIV f (x + z) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (x + z)) x :> y)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 742 | by (simp add: DERIV_iff field_simps) | 
| 21164 | 743 | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 744 | lemma DERIV_mirror: | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 745 | "(DERIV f (- x) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (- x::real) :: real) x :> - y)" | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 746 | by (simp add: deriv_def filterlim_at_split filterlim_at_left_to_right | 
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 747 | tendsto_minus_cancel_left field_simps conj_commute) | 
| 21164 | 748 | |
| 29975 | 749 | text {* Caratheodory formulation of derivative at a point *}
 | 
| 21164 | 750 | |
| 751 | lemma CARAT_DERIV: | |
| 51642 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 752 | "(DERIV f x :> l) \<longleftrightarrow> (\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> isCont g x \<and> g x = l)" | 
| 21164 | 753 | (is "?lhs = ?rhs") | 
| 754 | proof | |
| 755 | assume der: "DERIV f x :> l" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 756 | show "\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) \<and> isCont g x \<and> g x = l" | 
| 21164 | 757 | proof (intro exI conjI) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 758 | let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" | 
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23412diff
changeset | 759 | show "\<forall>z. f z - f x = ?g z * (z-x)" by simp | 
| 21164 | 760 | show "isCont ?g x" using der | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 761 | by (simp add: isCont_iff DERIV_iff cong: LIM_equal [rule_format]) | 
| 21164 | 762 | show "?g x = l" by simp | 
| 763 | qed | |
| 764 | next | |
| 765 | assume "?rhs" | |
| 766 | then obtain g where | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 767 | "(\<forall>z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast | 
| 21164 | 768 | thus "(DERIV f x :> l)" | 
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23412diff
changeset | 769 | by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong) | 
| 21164 | 770 | qed | 
| 771 | ||
| 31899 | 772 | text {*
 | 
| 773 | Let's do the standard proof, though theorem | |
| 774 |  @{text "LIM_mult2"} follows from a NS proof
 | |
| 775 | *} | |
| 21164 | 776 | |
| 29975 | 777 | subsection {* Local extrema *}
 | 
| 778 | ||
| 21164 | 779 | text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
 | 
| 780 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 781 | lemma DERIV_pos_inc_right: | 
| 21164 | 782 | fixes f :: "real => real" | 
| 783 | assumes der: "DERIV f x :> l" | |
| 784 | and l: "0 < l" | |
| 785 | shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x + h)" | |
| 786 | proof - | |
| 787 | from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] | |
| 788 | have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 789 | by simp | 
| 21164 | 790 | then obtain s | 
| 791 | where s: "0 < s" | |
| 792 | and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l" | |
| 793 | by auto | |
| 794 | thus ?thesis | |
| 795 | proof (intro exI conjI strip) | |
| 23441 | 796 | show "0<s" using s . | 
| 21164 | 797 | fix h::real | 
| 798 | assume "0 < h" "h < s" | |
| 799 | with all [of h] show "f x < f (x+h)" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 800 | proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm) | 
| 21164 | 801 | assume "~ (f (x+h) - f x) / h < l" and h: "0 < h" | 
| 802 | with l | |
| 803 | have "0 < (f (x+h) - f x) / h" by arith | |
| 804 | thus "f x < f (x+h)" | |
| 805 | by (simp add: pos_less_divide_eq h) | |
| 806 | qed | |
| 807 | qed | |
| 808 | qed | |
| 809 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 810 | lemma DERIV_neg_dec_left: | 
| 21164 | 811 | fixes f :: "real => real" | 
| 812 | assumes der: "DERIV f x :> l" | |
| 813 | and l: "l < 0" | |
| 814 | shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x-h)" | |
| 815 | proof - | |
| 816 | from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] | |
| 817 | have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 818 | by simp | 
| 21164 | 819 | then obtain s | 
| 820 | where s: "0 < s" | |
| 821 | and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l" | |
| 822 | by auto | |
| 823 | thus ?thesis | |
| 824 | proof (intro exI conjI strip) | |
| 23441 | 825 | show "0<s" using s . | 
| 21164 | 826 | fix h::real | 
| 827 | assume "0 < h" "h < s" | |
| 828 | with all [of "-h"] show "f x < f (x-h)" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 829 | proof (simp add: abs_if pos_less_divide_eq split add: split_if_asm) | 
| 21164 | 830 | assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h" | 
| 831 | with l | |
| 832 | have "0 < (f (x-h) - f x) / h" by arith | |
| 833 | thus "f x < f (x-h)" | |
| 834 | by (simp add: pos_less_divide_eq h) | |
| 835 | qed | |
| 836 | qed | |
| 837 | qed | |
| 838 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 839 | lemma DERIV_pos_inc_left: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 840 | fixes f :: "real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 841 | shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 842 | apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified]) | 
| 41368 | 843 | apply (auto simp add: DERIV_minus) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 844 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 845 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 846 | lemma DERIV_neg_dec_right: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 847 | fixes f :: "real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 848 | shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x) > f(x + h)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 849 | apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified]) | 
| 41368 | 850 | apply (auto simp add: DERIV_minus) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 851 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 852 | |
| 21164 | 853 | lemma DERIV_local_max: | 
| 854 | fixes f :: "real => real" | |
| 855 | assumes der: "DERIV f x :> l" | |
| 856 | and d: "0 < d" | |
| 857 | and le: "\<forall>y. \<bar>x-y\<bar> < d --> f(y) \<le> f(x)" | |
| 858 | shows "l = 0" | |
| 859 | proof (cases rule: linorder_cases [of l 0]) | |
| 23441 | 860 | case equal thus ?thesis . | 
| 21164 | 861 | next | 
| 862 | case less | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 863 | from DERIV_neg_dec_left [OF der less] | 
| 21164 | 864 | obtain d' where d': "0 < d'" | 
| 865 | and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x-h)" by blast | |
| 866 | from real_lbound_gt_zero [OF d d'] | |
| 867 | obtain e where "0 < e \<and> e < d \<and> e < d'" .. | |
| 868 | with lt le [THEN spec [where x="x-e"]] | |
| 869 | show ?thesis by (auto simp add: abs_if) | |
| 870 | next | |
| 871 | case greater | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 872 | from DERIV_pos_inc_right [OF der greater] | 
| 21164 | 873 | obtain d' where d': "0 < d'" | 
| 874 | and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" by blast | |
| 875 | from real_lbound_gt_zero [OF d d'] | |
| 876 | obtain e where "0 < e \<and> e < d \<and> e < d'" .. | |
| 877 | with lt le [THEN spec [where x="x+e"]] | |
| 878 | show ?thesis by (auto simp add: abs_if) | |
| 879 | qed | |
| 880 | ||
| 881 | ||
| 882 | text{*Similar theorem for a local minimum*}
 | |
| 883 | lemma DERIV_local_min: | |
| 884 | fixes f :: "real => real" | |
| 885 | shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0" | |
| 886 | by (drule DERIV_minus [THEN DERIV_local_max], auto) | |
| 887 | ||
| 888 | ||
| 889 | text{*In particular, if a function is locally flat*}
 | |
| 890 | lemma DERIV_local_const: | |
| 891 | fixes f :: "real => real" | |
| 892 | shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0" | |
| 893 | by (auto dest!: DERIV_local_max) | |
| 894 | ||
| 29975 | 895 | |
| 896 | subsection {* Rolle's Theorem *}
 | |
| 897 | ||
| 21164 | 898 | text{*Lemma about introducing open ball in open interval*}
 | 
| 899 | lemma lemma_interval_lt: | |
| 900 | "[| a < x; x < b |] | |
| 901 | ==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)" | |
| 27668 | 902 | |
| 22998 | 903 | apply (simp add: abs_less_iff) | 
| 21164 | 904 | apply (insert linorder_linear [of "x-a" "b-x"], safe) | 
| 905 | apply (rule_tac x = "x-a" in exI) | |
| 906 | apply (rule_tac [2] x = "b-x" in exI, auto) | |
| 907 | done | |
| 908 | ||
| 909 | lemma lemma_interval: "[| a < x; x < b |] ==> | |
| 910 | \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a \<le> y & y \<le> b)" | |
| 911 | apply (drule lemma_interval_lt, auto) | |
| 44921 | 912 | apply force | 
| 21164 | 913 | done | 
| 914 | ||
| 915 | text{*Rolle's Theorem.
 | |
| 916 |    If @{term f} is defined and continuous on the closed interval
 | |
| 917 |    @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
 | |
| 918 |    and @{term "f(a) = f(b)"},
 | |
| 919 |    then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
 | |
| 920 | theorem Rolle: | |
| 921 | assumes lt: "a < b" | |
| 922 | and eq: "f(a) = f(b)" | |
| 923 | and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x" | |
| 924 | and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 925 | shows "\<exists>z::real. a < z & z < b & DERIV f z :> 0" | 
| 21164 | 926 | proof - | 
| 927 | have le: "a \<le> b" using lt by simp | |
| 928 | from isCont_eq_Ub [OF le con] | |
| 929 | obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x" | |
| 930 | and alex: "a \<le> x" and xleb: "x \<le> b" | |
| 931 | by blast | |
| 932 | from isCont_eq_Lb [OF le con] | |
| 933 | obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z" | |
| 934 | and alex': "a \<le> x'" and x'leb: "x' \<le> b" | |
| 935 | by blast | |
| 936 | show ?thesis | |
| 937 | proof cases | |
| 938 | assume axb: "a < x & x < b" | |
| 939 |         --{*@{term f} attains its maximum within the interval*}
 | |
| 27668 | 940 | hence ax: "a<x" and xb: "x<b" by arith + | 
| 21164 | 941 | from lemma_interval [OF ax xb] | 
| 942 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 943 | by blast | |
| 944 | hence bound': "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> f y \<le> f x" using x_max | |
| 945 | by blast | |
| 946 | from differentiableD [OF dif [OF axb]] | |
| 947 | obtain l where der: "DERIV f x :> l" .. | |
| 948 | have "l=0" by (rule DERIV_local_max [OF der d bound']) | |
| 949 |         --{*the derivative at a local maximum is zero*}
 | |
| 950 | thus ?thesis using ax xb der by auto | |
| 951 | next | |
| 952 | assume notaxb: "~ (a < x & x < b)" | |
| 953 | hence xeqab: "x=a | x=b" using alex xleb by arith | |
| 954 | hence fb_eq_fx: "f b = f x" by (auto simp add: eq) | |
| 955 | show ?thesis | |
| 956 | proof cases | |
| 957 | assume ax'b: "a < x' & x' < b" | |
| 958 |         --{*@{term f} attains its minimum within the interval*}
 | |
| 27668 | 959 | hence ax': "a<x'" and x'b: "x'<b" by arith+ | 
| 21164 | 960 | from lemma_interval [OF ax' x'b] | 
| 961 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 962 | by blast | |
| 963 | hence bound': "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> f x' \<le> f y" using x'_min | |
| 964 | by blast | |
| 965 | from differentiableD [OF dif [OF ax'b]] | |
| 966 | obtain l where der: "DERIV f x' :> l" .. | |
| 967 | have "l=0" by (rule DERIV_local_min [OF der d bound']) | |
| 968 |         --{*the derivative at a local minimum is zero*}
 | |
| 969 | thus ?thesis using ax' x'b der by auto | |
| 970 | next | |
| 971 | assume notax'b: "~ (a < x' & x' < b)" | |
| 972 |         --{*@{term f} is constant througout the interval*}
 | |
| 973 | hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith | |
| 974 | hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) | |
| 975 | from dense [OF lt] | |
| 976 | obtain r where ar: "a < r" and rb: "r < b" by blast | |
| 977 | from lemma_interval [OF ar rb] | |
| 978 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 979 | by blast | |
| 980 | have eq_fb: "\<forall>z. a \<le> z --> z \<le> b --> f z = f b" | |
| 981 | proof (clarify) | |
| 982 | fix z::real | |
| 983 | assume az: "a \<le> z" and zb: "z \<le> b" | |
| 984 | show "f z = f b" | |
| 985 | proof (rule order_antisym) | |
| 986 | show "f z \<le> f b" by (simp add: fb_eq_fx x_max az zb) | |
| 987 | show "f b \<le> f z" by (simp add: fb_eq_fx' x'_min az zb) | |
| 988 | qed | |
| 989 | qed | |
| 990 | have bound': "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> f r = f y" | |
| 991 | proof (intro strip) | |
| 992 | fix y::real | |
| 993 | assume lt: "\<bar>r-y\<bar> < d" | |
| 994 | hence "f y = f b" by (simp add: eq_fb bound) | |
| 995 | thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) | |
| 996 | qed | |
| 997 | from differentiableD [OF dif [OF conjI [OF ar rb]]] | |
| 998 | obtain l where der: "DERIV f r :> l" .. | |
| 999 | have "l=0" by (rule DERIV_local_const [OF der d bound']) | |
| 1000 |         --{*the derivative of a constant function is zero*}
 | |
| 1001 | thus ?thesis using ar rb der by auto | |
| 1002 | qed | |
| 1003 | qed | |
| 1004 | qed | |
| 1005 | ||
| 1006 | ||
| 1007 | subsection{*Mean Value Theorem*}
 | |
| 1008 | ||
| 1009 | lemma lemma_MVT: | |
| 1010 | "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" | |
| 51481 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 hoelzl parents: 
51480diff
changeset | 1011 | by (cases "a = b") (simp_all add: field_simps) | 
| 21164 | 1012 | |
| 1013 | theorem MVT: | |
| 1014 | assumes lt: "a < b" | |
| 1015 | and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x" | |
| 1016 | and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1017 | shows "\<exists>l z::real. a < z & z < b & DERIV f z :> l & | 
| 21164 | 1018 | (f(b) - f(a) = (b-a) * l)" | 
| 1019 | proof - | |
| 1020 | let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" | |
| 44233 | 1021 | have contF: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?F x" | 
| 1022 | using con by (fast intro: isCont_intros) | |
| 21164 | 1023 | have difF: "\<forall>x. a < x \<and> x < b \<longrightarrow> ?F differentiable x" | 
| 1024 | proof (clarify) | |
| 1025 | fix x::real | |
| 1026 | assume ax: "a < x" and xb: "x < b" | |
| 1027 | from differentiableD [OF dif [OF conjI [OF ax xb]]] | |
| 1028 | obtain l where der: "DERIV f x :> l" .. | |
| 1029 | show "?F differentiable x" | |
| 1030 | by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], | |
| 1031 | blast intro: DERIV_diff DERIV_cmult_Id der) | |
| 1032 | qed | |
| 1033 | from Rolle [where f = ?F, OF lt lemma_MVT contF difF] | |
| 1034 | obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" | |
| 1035 | by blast | |
| 1036 | have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" | |
| 1037 | by (rule DERIV_cmult_Id) | |
| 1038 | hence derF: "DERIV (\<lambda>x. ?F x + (f b - f a) / (b - a) * x) z | |
| 1039 | :> 0 + (f b - f a) / (b - a)" | |
| 1040 | by (rule DERIV_add [OF der]) | |
| 1041 | show ?thesis | |
| 1042 | proof (intro exI conjI) | |
| 23441 | 1043 | show "a < z" using az . | 
| 1044 | show "z < b" using zb . | |
| 21164 | 1045 | show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) | 
| 1046 | show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp | |
| 1047 | qed | |
| 1048 | qed | |
| 1049 | ||
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1050 | lemma MVT2: | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1051 | "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |] | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1052 | ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1053 | apply (drule MVT) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1054 | apply (blast intro: DERIV_isCont) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1055 | apply (force dest: order_less_imp_le simp add: differentiable_def) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1056 | apply (blast dest: DERIV_unique order_less_imp_le) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1057 | done | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1058 | |
| 21164 | 1059 | |
| 1060 | text{*A function is constant if its derivative is 0 over an interval.*}
 | |
| 1061 | ||
| 1062 | lemma DERIV_isconst_end: | |
| 1063 | fixes f :: "real => real" | |
| 1064 | shows "[| a < b; | |
| 1065 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1066 | \<forall>x. a < x & x < b --> DERIV f x :> 0 |] | |
| 1067 | ==> f b = f a" | |
| 1068 | apply (drule MVT, assumption) | |
| 1069 | apply (blast intro: differentiableI) | |
| 1070 | apply (auto dest!: DERIV_unique simp add: diff_eq_eq) | |
| 1071 | done | |
| 1072 | ||
| 1073 | lemma DERIV_isconst1: | |
| 1074 | fixes f :: "real => real" | |
| 1075 | shows "[| a < b; | |
| 1076 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1077 | \<forall>x. a < x & x < b --> DERIV f x :> 0 |] | |
| 1078 | ==> \<forall>x. a \<le> x & x \<le> b --> f x = f a" | |
| 1079 | apply safe | |
| 1080 | apply (drule_tac x = a in order_le_imp_less_or_eq, safe) | |
| 1081 | apply (drule_tac b = x in DERIV_isconst_end, auto) | |
| 1082 | done | |
| 1083 | ||
| 1084 | lemma DERIV_isconst2: | |
| 1085 | fixes f :: "real => real" | |
| 1086 | shows "[| a < b; | |
| 1087 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1088 | \<forall>x. a < x & x < b --> DERIV f x :> 0; | |
| 1089 | a \<le> x; x \<le> b |] | |
| 1090 | ==> f x = f a" | |
| 1091 | apply (blast dest: DERIV_isconst1) | |
| 1092 | done | |
| 1093 | ||
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1094 | lemma DERIV_isconst3: fixes a b x y :: real | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1095 |   assumes "a < b" and "x \<in> {a <..< b}" and "y \<in> {a <..< b}"
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1096 |   assumes derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1097 | shows "f x = f y" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1098 | proof (cases "x = y") | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1099 | case False | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1100 | let ?a = "min x y" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1101 | let ?b = "max x y" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1102 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1103 | have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1104 | proof (rule allI, rule impI) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1105 | fix z :: real assume "?a \<le> z \<and> z \<le> ?b" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1106 |     hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1107 |     hence "z \<in> {a<..<b}" by auto
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1108 | thus "DERIV f z :> 0" by (rule derivable) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1109 | qed | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1110 | hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1111 | and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1112 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1113 | have "?a < ?b" using `x \<noteq> y` by auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1114 | from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1115 | show ?thesis by auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1116 | qed auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1117 | |
| 21164 | 1118 | lemma DERIV_isconst_all: | 
| 1119 | fixes f :: "real => real" | |
| 1120 | shows "\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)" | |
| 1121 | apply (rule linorder_cases [of x y]) | |
| 1122 | apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ | |
| 1123 | done | |
| 1124 | ||
| 1125 | lemma DERIV_const_ratio_const: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1126 | fixes f :: "real => real" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1127 | shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" | 
| 21164 | 1128 | apply (rule linorder_cases [of a b], auto) | 
| 1129 | apply (drule_tac [!] f = f in MVT) | |
| 1130 | apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53381diff
changeset | 1131 | apply (auto dest: DERIV_unique simp add: ring_distribs) | 
| 21164 | 1132 | done | 
| 1133 | ||
| 1134 | lemma DERIV_const_ratio_const2: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1135 | fixes f :: "real => real" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1136 | shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" | 
| 21164 | 1137 | apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) | 
| 1138 | apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) | |
| 1139 | done | |
| 1140 | ||
| 1141 | lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" | |
| 1142 | by (simp) | |
| 1143 | ||
| 1144 | lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" | |
| 1145 | by (simp) | |
| 1146 | ||
| 1147 | text{*Gallileo's "trick": average velocity = av. of end velocities*}
 | |
| 1148 | ||
| 1149 | lemma DERIV_const_average: | |
| 1150 | fixes v :: "real => real" | |
| 1151 | assumes neq: "a \<noteq> (b::real)" | |
| 1152 | and der: "\<forall>x. DERIV v x :> k" | |
| 1153 | shows "v ((a + b)/2) = (v a + v b)/2" | |
| 1154 | proof (cases rule: linorder_cases [of a b]) | |
| 1155 | case equal with neq show ?thesis by simp | |
| 1156 | next | |
| 1157 | case less | |
| 1158 | have "(v b - v a) / (b - a) = k" | |
| 1159 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 1160 | hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp | |
| 1161 | moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" | |
| 1162 | by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) | |
| 1163 | ultimately show ?thesis using neq by force | |
| 1164 | next | |
| 1165 | case greater | |
| 1166 | have "(v b - v a) / (b - a) = k" | |
| 1167 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 1168 | hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp | |
| 1169 | moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" | |
| 1170 | by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) | |
| 1171 | ultimately show ?thesis using neq by (force simp add: add_commute) | |
| 1172 | qed | |
| 1173 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1174 | (* A function with positive derivative is increasing. | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1175 | A simple proof using the MVT, by Jeremy Avigad. And variants. | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1176 | *) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1177 | lemma DERIV_pos_imp_increasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1178 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1179 | assumes "a < b" and "\<forall>x. a \<le> x & x \<le> b --> (EX y. DERIV f x :> y & y > 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1180 | shows "f a < f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1181 | proof (rule ccontr) | 
| 41550 | 1182 | assume f: "~ f a < f b" | 
| 33690 | 1183 | have "EX l z. a < z & z < b & DERIV f z :> l | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1184 | & f b - f a = (b - a) * l" | 
| 33690 | 1185 | apply (rule MVT) | 
| 1186 | using assms | |
| 1187 | apply auto | |
| 1188 | apply (metis DERIV_isCont) | |
| 36777 
be5461582d0f
avoid using real-specific versions of generic lemmas
 huffman parents: 
35216diff
changeset | 1189 | apply (metis differentiableI less_le) | 
| 33690 | 1190 | done | 
| 41550 | 1191 | then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1192 | and "f b - f a = (b - a) * l" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1193 | by auto | 
| 41550 | 1194 | with assms f have "~(l > 0)" | 
| 36777 
be5461582d0f
avoid using real-specific versions of generic lemmas
 huffman parents: 
35216diff
changeset | 1195 | by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) | 
| 41550 | 1196 | with assms z show False | 
| 36777 
be5461582d0f
avoid using real-specific versions of generic lemmas
 huffman parents: 
35216diff
changeset | 1197 | by (metis DERIV_unique less_le) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1198 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1199 | |
| 45791 | 1200 | lemma DERIV_nonneg_imp_nondecreasing: | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1201 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1202 | assumes "a \<le> b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1203 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<ge> 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1204 | shows "f a \<le> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1205 | proof (rule ccontr, cases "a = b") | 
| 41550 | 1206 | assume "~ f a \<le> f b" and "a = b" | 
| 1207 | then show False by auto | |
| 37891 | 1208 | next | 
| 1209 | assume A: "~ f a \<le> f b" | |
| 1210 | assume B: "a ~= b" | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1211 | with assms have "EX l z. a < z & z < b & DERIV f z :> l | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1212 | & f b - f a = (b - a) * l" | 
| 33690 | 1213 | apply - | 
| 1214 | apply (rule MVT) | |
| 1215 | apply auto | |
| 1216 | apply (metis DERIV_isCont) | |
| 36777 
be5461582d0f
avoid using real-specific versions of generic lemmas
 huffman parents: 
35216diff
changeset | 1217 | apply (metis differentiableI less_le) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1218 | done | 
| 41550 | 1219 | then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" | 
| 37891 | 1220 | and C: "f b - f a = (b - a) * l" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1221 | by auto | 
| 37891 | 1222 | with A have "a < b" "f b < f a" by auto | 
| 1223 | with C have "\<not> l \<ge> 0" by (auto simp add: not_le algebra_simps) | |
| 45051 
c478d1876371
discontinued legacy theorem names from RealDef.thy
 huffman parents: 
44921diff
changeset | 1224 | (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) | 
| 41550 | 1225 | with assms z show False | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1226 | by (metis DERIV_unique order_less_imp_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1227 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1228 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1229 | lemma DERIV_neg_imp_decreasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1230 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1231 | assumes "a < b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1232 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y < 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1233 | shows "f a > f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1234 | proof - | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1235 | have "(%x. -f x) a < (%x. -f x) b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1236 | apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"]) | 
| 33690 | 1237 | using assms | 
| 1238 | apply auto | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1239 | apply (metis DERIV_minus neg_0_less_iff_less) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1240 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1241 | thus ?thesis | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1242 | by simp | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1243 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1244 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1245 | lemma DERIV_nonpos_imp_nonincreasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1246 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1247 | assumes "a \<le> b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1248 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<le> 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1249 | shows "f a \<ge> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1250 | proof - | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1251 | have "(%x. -f x) a \<le> (%x. -f x) b" | 
| 45791 | 1252 | apply (rule DERIV_nonneg_imp_nondecreasing [of a b "%x. -f x"]) | 
| 33690 | 1253 | using assms | 
| 1254 | apply auto | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1255 | apply (metis DERIV_minus neg_0_le_iff_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1256 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1257 | thus ?thesis | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1258 | by simp | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1259 | qed | 
| 21164 | 1260 | |
| 23041 | 1261 | text {* Derivative of inverse function *}
 | 
| 1262 | ||
| 1263 | lemma DERIV_inverse_function: | |
| 1264 | fixes f g :: "real \<Rightarrow> real" | |
| 1265 | assumes der: "DERIV f (g x) :> D" | |
| 1266 | assumes neq: "D \<noteq> 0" | |
| 23044 | 1267 | assumes a: "a < x" and b: "x < b" | 
| 1268 | assumes inj: "\<forall>y. a < y \<and> y < b \<longrightarrow> f (g y) = y" | |
| 23041 | 1269 | assumes cont: "isCont g x" | 
| 1270 | shows "DERIV g x :> inverse D" | |
| 1271 | unfolding DERIV_iff2 | |
| 23044 | 1272 | proof (rule LIM_equal2) | 
| 1273 | show "0 < min (x - a) (b - x)" | |
| 27668 | 1274 | using a b by arith | 
| 23044 | 1275 | next | 
| 23041 | 1276 | fix y | 
| 23044 | 1277 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 27668 | 1278 | hence "a < y" and "y < b" | 
| 23044 | 1279 | by (simp_all add: abs_less_iff) | 
| 23041 | 1280 | thus "(g y - g x) / (y - x) = | 
| 1281 | inverse ((f (g y) - x) / (g y - g x))" | |
| 1282 | by (simp add: inj) | |
| 1283 | next | |
| 1284 | have "(\<lambda>z. (f z - f (g x)) / (z - g x)) -- g x --> D" | |
| 1285 | by (rule der [unfolded DERIV_iff2]) | |
| 1286 | hence 1: "(\<lambda>z. (f z - x) / (z - g x)) -- g x --> D" | |
| 23044 | 1287 | using inj a b by simp | 
| 23041 | 1288 | have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x" | 
| 1289 | proof (safe intro!: exI) | |
| 23044 | 1290 | show "0 < min (x - a) (b - x)" | 
| 1291 | using a b by simp | |
| 23041 | 1292 | next | 
| 1293 | fix y | |
| 23044 | 1294 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 1295 | hence y: "a < y" "y < b" | |
| 1296 | by (simp_all add: abs_less_iff) | |
| 23041 | 1297 | assume "g y = g x" | 
| 1298 | hence "f (g y) = f (g x)" by simp | |
| 23044 | 1299 | hence "y = x" using inj y a b by simp | 
| 23041 | 1300 | also assume "y \<noteq> x" | 
| 1301 | finally show False by simp | |
| 1302 | qed | |
| 1303 | have "(\<lambda>y. (f (g y) - x) / (g y - g x)) -- x --> D" | |
| 1304 | using cont 1 2 by (rule isCont_LIM_compose2) | |
| 1305 | thus "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x))) | |
| 1306 | -- x --> inverse D" | |
| 44568 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44317diff
changeset | 1307 | using neq by (rule tendsto_inverse) | 
| 23041 | 1308 | qed | 
| 1309 | ||
| 29975 | 1310 | subsection {* Generalized Mean Value Theorem *}
 | 
| 1311 | ||
| 21164 | 1312 | theorem GMVT: | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1313 | fixes a b :: real | 
| 21164 | 1314 | assumes alb: "a < b" | 
| 41550 | 1315 | and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x" | 
| 1316 | and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable x" | |
| 1317 | and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x" | |
| 1318 | and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable x" | |
| 53381 | 1319 | shows "\<exists>g'c f'c c. | 
| 1320 | DERIV g c :> g'c \<and> DERIV f c :> f'c \<and> a < c \<and> c < b \<and> ((f b - f a) * g'c) = ((g b - g a) * f'c)" | |
| 21164 | 1321 | proof - | 
| 1322 | let ?h = "\<lambda>x. (f b - f a)*(g x) - (g b - g a)*(f x)" | |
| 41550 | 1323 | from assms have "a < b" by simp | 
| 21164 | 1324 | moreover have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?h x" | 
| 44233 | 1325 | using fc gc by simp | 
| 1326 | moreover have "\<forall>x. a < x \<and> x < b \<longrightarrow> ?h differentiable x" | |
| 1327 | using fd gd by simp | |
| 21164 | 1328 | ultimately have "\<exists>l z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" by (rule MVT) | 
| 1329 | then obtain l where ldef: "\<exists>z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 1330 | then obtain c where cdef: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 1331 | ||
| 1332 | from cdef have cint: "a < c \<and> c < b" by auto | |
| 1333 | with gd have "g differentiable c" by simp | |
| 1334 | hence "\<exists>D. DERIV g c :> D" by (rule differentiableD) | |
| 1335 | then obtain g'c where g'cdef: "DERIV g c :> g'c" .. | |
| 1336 | ||
| 1337 | from cdef have "a < c \<and> c < b" by auto | |
| 1338 | with fd have "f differentiable c" by simp | |
| 1339 | hence "\<exists>D. DERIV f c :> D" by (rule differentiableD) | |
| 1340 | then obtain f'c where f'cdef: "DERIV f c :> f'c" .. | |
| 1341 | ||
| 1342 | from cdef have "DERIV ?h c :> l" by auto | |
| 41368 | 1343 | moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" | 
| 1344 | using g'cdef f'cdef by (auto intro!: DERIV_intros) | |
| 21164 | 1345 | ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) | 
| 1346 | ||
| 1347 |   {
 | |
| 1348 | from cdef have "?h b - ?h a = (b - a) * l" by auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 1349 | also from leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp | 
| 21164 | 1350 | finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp | 
| 1351 | } | |
| 1352 | moreover | |
| 1353 |   {
 | |
| 1354 | have "?h b - ?h a = | |
| 1355 | ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - | |
| 1356 | ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" | |
| 29667 | 1357 | by (simp add: algebra_simps) | 
| 21164 | 1358 | hence "?h b - ?h a = 0" by auto | 
| 1359 | } | |
| 1360 | ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto | |
| 1361 | with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp | |
| 1362 | hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp | |
| 1363 | hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) | |
| 1364 | ||
| 1365 | with g'cdef f'cdef cint show ?thesis by auto | |
| 1366 | qed | |
| 1367 | ||
| 50327 | 1368 | lemma GMVT': | 
| 1369 | fixes f g :: "real \<Rightarrow> real" | |
| 1370 | assumes "a < b" | |
| 1371 | assumes isCont_f: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont f z" | |
| 1372 | assumes isCont_g: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont g z" | |
| 1373 | assumes DERIV_g: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV g z :> (g' z)" | |
| 1374 | assumes DERIV_f: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV f z :> (f' z)" | |
| 1375 | shows "\<exists>c. a < c \<and> c < b \<and> (f b - f a) * g' c = (g b - g a) * f' c" | |
| 1376 | proof - | |
| 1377 | have "\<exists>g'c f'c c. DERIV g c :> g'c \<and> DERIV f c :> f'c \<and> | |
| 1378 | a < c \<and> c < b \<and> (f b - f a) * g'c = (g b - g a) * f'c" | |
| 1379 | using assms by (intro GMVT) (force simp: differentiable_def)+ | |
| 1380 | then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" | |
| 1381 | using DERIV_f DERIV_g by (force dest: DERIV_unique) | |
| 1382 | then show ?thesis | |
| 1383 | by auto | |
| 1384 | qed | |
| 1385 | ||
| 51529 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 1386 | |
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 1387 | subsection {* L'Hopitals rule *}
 | 
| 
2d2f59e6055a
move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
 hoelzl parents: 
51526diff
changeset | 1388 | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1389 | lemma isCont_If_ge: | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1390 | fixes a :: "'a :: linorder_topology" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1391 | shows "continuous (at_left a) g \<Longrightarrow> (f ---> g a) (at_right a) \<Longrightarrow> isCont (\<lambda>x. if x \<le> a then g x else f x) a" | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1392 | unfolding isCont_def continuous_within | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1393 | apply (intro filterlim_split_at) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1394 | apply (subst filterlim_cong[OF refl refl, where g=g]) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1395 | apply (simp_all add: eventually_at_filter less_le) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1396 | apply (subst filterlim_cong[OF refl refl, where g=f]) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1397 | apply (simp_all add: eventually_at_filter less_le) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1398 | done | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1399 | |
| 50327 | 1400 | lemma lhopital_right_0: | 
| 50329 | 1401 | fixes f0 g0 :: "real \<Rightarrow> real" | 
| 1402 | assumes f_0: "(f0 ---> 0) (at_right 0)" | |
| 1403 | assumes g_0: "(g0 ---> 0) (at_right 0)" | |
| 50327 | 1404 | assumes ev: | 
| 50329 | 1405 | "eventually (\<lambda>x. g0 x \<noteq> 0) (at_right 0)" | 
| 50327 | 1406 | "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)" | 
| 50329 | 1407 | "eventually (\<lambda>x. DERIV f0 x :> f' x) (at_right 0)" | 
| 1408 | "eventually (\<lambda>x. DERIV g0 x :> g' x) (at_right 0)" | |
| 50327 | 1409 | assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)" | 
| 50329 | 1410 | shows "((\<lambda> x. f0 x / g0 x) ---> x) (at_right 0)" | 
| 50327 | 1411 | proof - | 
| 50329 | 1412 | def f \<equiv> "\<lambda>x. if x \<le> 0 then 0 else f0 x" | 
| 1413 | then have "f 0 = 0" by simp | |
| 1414 | ||
| 1415 | def g \<equiv> "\<lambda>x. if x \<le> 0 then 0 else g0 x" | |
| 1416 | then have "g 0 = 0" by simp | |
| 1417 | ||
| 1418 | have "eventually (\<lambda>x. g0 x \<noteq> 0 \<and> g' x \<noteq> 0 \<and> | |
| 1419 | DERIV f0 x :> (f' x) \<and> DERIV g0 x :> (g' x)) (at_right 0)" | |
| 1420 | using ev by eventually_elim auto | |
| 1421 | then obtain a where [arith]: "0 < a" | |
| 1422 | and g0_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g0 x \<noteq> 0" | |
| 50327 | 1423 | and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0" | 
| 50329 | 1424 | and f0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV f0 x :> (f' x)" | 
| 1425 | and g0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV g0 x :> (g' x)" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1426 | unfolding eventually_at eventually_at by (auto simp: dist_real_def) | 
| 50327 | 1427 | |
| 50329 | 1428 | have g_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g x \<noteq> 0" | 
| 1429 | using g0_neq_0 by (simp add: g_def) | |
| 1430 | ||
| 1431 |   { fix x assume x: "0 < x" "x < a" then have "DERIV f x :> (f' x)"
 | |
| 1432 | by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) | |
| 1433 | (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) } | |
| 1434 | note f = this | |
| 1435 | ||
| 1436 |   { fix x assume x: "0 < x" "x < a" then have "DERIV g x :> (g' x)"
 | |
| 1437 | by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) | |
| 1438 | (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) } | |
| 1439 | note g = this | |
| 1440 | ||
| 1441 | have "isCont f 0" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1442 | unfolding f_def by (intro isCont_If_ge f_0 continuous_const) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1443 | |
| 50329 | 1444 | have "isCont g 0" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1445 | unfolding g_def by (intro isCont_If_ge g_0 continuous_const) | 
| 50329 | 1446 | |
| 50327 | 1447 |   have "\<exists>\<zeta>. \<forall>x\<in>{0 <..< a}. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)"
 | 
| 1448 | proof (rule bchoice, rule) | |
| 1449 |     fix x assume "x \<in> {0 <..< a}"
 | |
| 1450 | then have x[arith]: "0 < x" "x < a" by auto | |
| 1451 | with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x" | |
| 1452 | by auto | |
| 50328 | 1453 | have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont f x" | 
| 1454 | using `isCont f 0` f by (auto intro: DERIV_isCont simp: le_less) | |
| 1455 | moreover have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont g x" | |
| 1456 | using `isCont g 0` g by (auto intro: DERIV_isCont simp: le_less) | |
| 1457 | ultimately have "\<exists>c. 0 < c \<and> c < x \<and> (f x - f 0) * g' c = (g x - g 0) * f' c" | |
| 1458 | using f g `x < a` by (intro GMVT') auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 1459 | then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 1460 | by blast | 
| 50327 | 1461 | moreover | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 1462 | from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" | 
| 50327 | 1463 | by (simp add: field_simps) | 
| 1464 | ultimately show "\<exists>y. 0 < y \<and> y < x \<and> f x / g x = f' y / g' y" | |
| 1465 | using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c]) | |
| 1466 | qed | |
| 53381 | 1467 |   then obtain \<zeta> where "\<forall>x\<in>{0 <..< a}. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)" ..
 | 
| 50327 | 1468 | then have \<zeta>: "eventually (\<lambda>x. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)) (at_right 0)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1469 | unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) | 
| 50327 | 1470 | moreover | 
| 1471 | from \<zeta> have "eventually (\<lambda>x. norm (\<zeta> x) \<le> x) (at_right 0)" | |
| 1472 | by eventually_elim auto | |
| 1473 | then have "((\<lambda>x. norm (\<zeta> x)) ---> 0) (at_right 0)" | |
| 1474 | by (rule_tac real_tendsto_sandwich[where f="\<lambda>x. 0" and h="\<lambda>x. x"]) | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1475 | (auto intro: tendsto_const tendsto_ident_at) | 
| 50327 | 1476 | then have "(\<zeta> ---> 0) (at_right 0)" | 
| 1477 | by (rule tendsto_norm_zero_cancel) | |
| 1478 | with \<zeta> have "filterlim \<zeta> (at_right 0) (at_right 0)" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1479 | by (auto elim!: eventually_elim1 simp: filterlim_at) | 
| 50327 | 1480 | from this lim have "((\<lambda>t. f' (\<zeta> t) / g' (\<zeta> t)) ---> x) (at_right 0)" | 
| 1481 | by (rule_tac filterlim_compose[of _ _ _ \<zeta>]) | |
| 50329 | 1482 | ultimately have "((\<lambda>t. f t / g t) ---> x) (at_right 0)" (is ?P) | 
| 50328 | 1483 | by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) | 
| 1484 | (auto elim: eventually_elim1) | |
| 50329 | 1485 | also have "?P \<longleftrightarrow> ?thesis" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1486 | by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) | 
| 50329 | 1487 | finally show ?thesis . | 
| 50327 | 1488 | qed | 
| 1489 | ||
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1490 | lemma lhopital_right: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1491 | "((f::real \<Rightarrow> real) ---> 0) (at_right x) \<Longrightarrow> (g ---> 0) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1492 | eventually (\<lambda>x. g x \<noteq> 0) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1493 | eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1494 | eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1495 | eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1496 | ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1497 | ((\<lambda> x. f x / g x) ---> y) (at_right x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1498 | unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1499 | by (rule lhopital_right_0) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1500 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1501 | lemma lhopital_left: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1502 | "((f::real \<Rightarrow> real) ---> 0) (at_left x) \<Longrightarrow> (g ---> 0) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1503 | eventually (\<lambda>x. g x \<noteq> 0) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1504 | eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1505 | eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1506 | eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1507 | ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1508 | ((\<lambda> x. f x / g x) ---> y) (at_left x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1509 | unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1510 | by (rule lhopital_right[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1511 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1512 | lemma lhopital: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1513 | "((f::real \<Rightarrow> real) ---> 0) (at x) \<Longrightarrow> (g ---> 0) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1514 | eventually (\<lambda>x. g x \<noteq> 0) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1515 | eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1516 | eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1517 | eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1518 | ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1519 | ((\<lambda> x. f x / g x) ---> y) (at x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1520 | unfolding eventually_at_split filterlim_at_split | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1521 | by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1522 | |
| 50327 | 1523 | lemma lhopital_right_0_at_top: | 
| 1524 | fixes f g :: "real \<Rightarrow> real" | |
| 1525 | assumes g_0: "LIM x at_right 0. g x :> at_top" | |
| 1526 | assumes ev: | |
| 1527 | "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)" | |
| 1528 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_right 0)" | |
| 1529 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_right 0)" | |
| 1530 | assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)" | |
| 1531 | shows "((\<lambda> x. f x / g x) ---> x) (at_right 0)" | |
| 1532 | unfolding tendsto_iff | |
| 1533 | proof safe | |
| 1534 | fix e :: real assume "0 < e" | |
| 1535 | ||
| 1536 | with lim[unfolded tendsto_iff, rule_format, of "e / 4"] | |
| 1537 | have "eventually (\<lambda>t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp | |
| 1538 | from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] | |
| 1539 | obtain a where [arith]: "0 < a" | |
| 1540 | and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0" | |
| 1541 | and f0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV f x :> (f' x)" | |
| 1542 | and g0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV g x :> (g' x)" | |
| 1543 | and Df: "\<And>t. 0 < t \<Longrightarrow> t < a \<Longrightarrow> dist (f' t / g' t) x < e / 4" | |
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1544 | unfolding eventually_at_le by (auto simp: dist_real_def) | 
| 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1545 | |
| 50327 | 1546 | |
| 1547 | from Df have | |
| 50328 | 1548 | "eventually (\<lambda>t. t < a) (at_right 0)" "eventually (\<lambda>t::real. 0 < t) (at_right 0)" | 
| 51641 
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
 hoelzl parents: 
51529diff
changeset | 1549 | unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) | 
| 50327 | 1550 | |
| 1551 | moreover | |
| 50328 | 1552 | have "eventually (\<lambda>t. 0 < g t) (at_right 0)" "eventually (\<lambda>t. g a < g t) (at_right 0)" | 
| 50346 
a75c6429c3c3
add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
 hoelzl parents: 
50331diff
changeset | 1553 | using g_0 by (auto elim: eventually_elim1 simp: filterlim_at_top_dense) | 
| 50327 | 1554 | |
| 1555 | moreover | |
| 1556 | have inv_g: "((\<lambda>x. inverse (g x)) ---> 0) (at_right 0)" | |
| 1557 | using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] | |
| 1558 | by (rule filterlim_compose) | |
| 1559 | then have "((\<lambda>x. norm (1 - g a * inverse (g x))) ---> norm (1 - g a * 0)) (at_right 0)" | |
| 1560 | by (intro tendsto_intros) | |
| 1561 | then have "((\<lambda>x. norm (1 - g a / g x)) ---> 1) (at_right 0)" | |
| 1562 | by (simp add: inverse_eq_divide) | |
| 1563 | from this[unfolded tendsto_iff, rule_format, of 1] | |
| 1564 | have "eventually (\<lambda>x. norm (1 - g a / g x) < 2) (at_right 0)" | |
| 1565 | by (auto elim!: eventually_elim1 simp: dist_real_def) | |
| 1566 | ||
| 1567 | moreover | |
| 1568 | from inv_g have "((\<lambda>t. norm ((f a - x * g a) * inverse (g t))) ---> norm ((f a - x * g a) * 0)) (at_right 0)" | |
| 1569 | by (intro tendsto_intros) | |
| 1570 | then have "((\<lambda>t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)" | |
| 1571 | by (simp add: inverse_eq_divide) | |
| 1572 | from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e` | |
| 1573 | have "eventually (\<lambda>t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" | |
| 1574 | by (auto simp: dist_real_def) | |
| 1575 | ||
| 1576 | ultimately show "eventually (\<lambda>t. dist (f t / g t) x < e) (at_right 0)" | |
| 1577 | proof eventually_elim | |
| 1578 | fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" | |
| 1579 | assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" | |
| 1580 | ||
| 1581 | have "\<exists>y. t < y \<and> y < a \<and> (g a - g t) * f' y = (f a - f t) * g' y" | |
| 1582 | using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ | |
| 53381 | 1583 | then obtain y where [arith]: "t < y" "y < a" | 
| 1584 | and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" | |
| 1585 | by blast | |
| 1586 | from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" | |
| 50327 | 1587 | using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps) | 
| 1588 | ||
| 1589 | 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" | |
| 1590 | by (simp add: field_simps) | |
| 1591 | have "norm (f t / g t - x) \<le> | |
| 1592 | norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" | |
| 1593 | unfolding * by (rule norm_triangle_ineq) | |
| 1594 | also have "\<dots> = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" | |
| 1595 | by (simp add: abs_mult D_eq dist_real_def) | |
| 1596 | also have "\<dots> < (e / 4) * 2 + e / 2" | |
| 1597 | using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto | |
| 1598 | finally show "dist (f t / g t) x < e" | |
| 1599 | by (simp add: dist_real_def) | |
| 1600 | qed | |
| 1601 | qed | |
| 1602 | ||
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1603 | lemma lhopital_right_at_top: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1604 | "LIM x at_right x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1605 | eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1606 | eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1607 | eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1608 | ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1609 | ((\<lambda> x. f x / g x) ---> y) (at_right x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1610 | unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1611 | by (rule lhopital_right_0_at_top) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1612 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1613 | lemma lhopital_left_at_top: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1614 | "LIM x at_left x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1615 | eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1616 | eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1617 | eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1618 | ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1619 | ((\<lambda> x. f x / g x) ---> y) (at_left x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1620 | unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1621 | by (rule lhopital_right_at_top[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1622 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1623 | lemma lhopital_at_top: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1624 | "LIM x at x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1625 | eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1626 | eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1627 | eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1628 | ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow> | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1629 | ((\<lambda> x. f x / g x) ---> y) (at x)" | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1630 | unfolding eventually_at_split filterlim_at_split | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1631 | by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 1632 | |
| 50347 | 1633 | lemma lhospital_at_top_at_top: | 
| 1634 | fixes f g :: "real \<Rightarrow> real" | |
| 1635 | assumes g_0: "LIM x at_top. g x :> at_top" | |
| 1636 | assumes g': "eventually (\<lambda>x. g' x \<noteq> 0) at_top" | |
| 1637 | assumes Df: "eventually (\<lambda>x. DERIV f x :> f' x) at_top" | |
| 1638 | assumes Dg: "eventually (\<lambda>x. DERIV g x :> g' x) at_top" | |
| 1639 | assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) at_top" | |
| 1640 | shows "((\<lambda> x. f x / g x) ---> x) at_top" | |
| 1641 | unfolding filterlim_at_top_to_right | |
| 1642 | proof (rule lhopital_right_0_at_top) | |
| 1643 | let ?F = "\<lambda>x. f (inverse x)" | |
| 1644 | let ?G = "\<lambda>x. g (inverse x)" | |
| 1645 | let ?R = "at_right (0::real)" | |
| 1646 | let ?D = "\<lambda>f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" | |
| 1647 | ||
| 1648 | show "LIM x ?R. ?G x :> at_top" | |
| 1649 | using g_0 unfolding filterlim_at_top_to_right . | |
| 1650 | ||
| 1651 | show "eventually (\<lambda>x. DERIV ?G x :> ?D g' x) ?R" | |
| 1652 | unfolding eventually_at_right_to_top | |
| 1653 | using Dg eventually_ge_at_top[where c="1::real"] | |
| 1654 | apply eventually_elim | |
| 1655 | apply (rule DERIV_cong) | |
| 1656 | apply (rule DERIV_chain'[where f=inverse]) | |
| 1657 | apply (auto intro!: DERIV_inverse) | |
| 1658 | done | |
| 1659 | ||
| 1660 | show "eventually (\<lambda>x. DERIV ?F x :> ?D f' x) ?R" | |
| 1661 | unfolding eventually_at_right_to_top | |
| 1662 | using Df eventually_ge_at_top[where c="1::real"] | |
| 1663 | apply eventually_elim | |
| 1664 | apply (rule DERIV_cong) | |
| 1665 | apply (rule DERIV_chain'[where f=inverse]) | |
| 1666 | apply (auto intro!: DERIV_inverse) | |
| 1667 | done | |
| 1668 | ||
| 1669 | show "eventually (\<lambda>x. ?D g' x \<noteq> 0) ?R" | |
| 1670 | unfolding eventually_at_right_to_top | |
| 1671 | using g' eventually_ge_at_top[where c="1::real"] | |
| 1672 | by eventually_elim auto | |
| 1673 | ||
| 1674 | show "((\<lambda>x. ?D f' x / ?D g' x) ---> x) ?R" | |
| 1675 | unfolding filterlim_at_right_to_top | |
| 1676 | apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) | |
| 1677 | using eventually_ge_at_top[where c="1::real"] | |
| 1678 | by eventually_elim simp | |
| 1679 | qed | |
| 1680 | ||
| 21164 | 1681 | end |