| author | wenzelm | 
| Fri, 15 Oct 2021 01:44:52 +0200 | |
| changeset 74519 | fc65e39ca170 | 
| parent 74007 | df976eefcba0 | 
| child 74878 | 0263787a06b4 | 
| permissions | -rw-r--r-- | 
| 63558 | 1 | (* Title: HOL/Deriv.thy | 
| 2 | Author: Jacques D. Fleuriot, University of Cambridge, 1998 | |
| 3 | Author: Brian Huffman | |
| 4 | Author: Lawrence C Paulson, 2004 | |
| 5 | Author: Benjamin Porter, 2005 | |
| 21164 | 6 | *) | 
| 7 | ||
| 63558 | 8 | section \<open>Differentiation\<close> | 
| 21164 | 9 | |
| 10 | theory Deriv | |
| 63558 | 11 | imports Limits | 
| 21164 | 12 | begin | 
| 13 | ||
| 60758 | 14 | subsection \<open>Frechet derivative\<close> | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 15 | |
| 63558 | 16 | definition has_derivative :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow>
 | 
| 17 |     ('a \<Rightarrow> 'b) \<Rightarrow> 'a filter \<Rightarrow> bool"  (infix "(has'_derivative)" 50)
 | |
| 18 | where "(f has_derivative f') F \<longleftrightarrow> | |
| 19 | bounded_linear f' \<and> | |
| 20 | ((\<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))) \<longlongrightarrow> 0) 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 | 21 | |
| 60758 | 22 | text \<open> | 
| 69593 | 23 | Usually the filter \<^term>\<open>F\<close> is \<^term>\<open>at x within s\<close>. \<^term>\<open>(f has_derivative D) | 
| 24 | (at x within s)\<close> means: \<^term>\<open>D\<close> is the derivative of function \<^term>\<open>f\<close> at point \<^term>\<open>x\<close> | |
| 25 | within the set \<^term>\<open>s\<close>. Where \<^term>\<open>s\<close> is used to express left or right sided derivatives. In | |
| 26 | most cases \<^term>\<open>s\<close> is either a variable or \<^term>\<open>UNIV\<close>. | |
| 60758 | 27 | \<close> | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 28 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 29 | text \<open>These are the only cases we'll care about, probably.\<close> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 30 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 31 | lemma has_derivative_within: "(f has_derivative f') (at x within s) \<longleftrightarrow> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 32 | bounded_linear f' \<and> ((\<lambda>y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 33 | unfolding has_derivative_def tendsto_iff | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 34 | by (subst eventually_Lim_ident_at) (auto simp add: field_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 35 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 36 | lemma has_derivative_eq_rhs: "(f has_derivative f') F \<Longrightarrow> f' = g' \<Longrightarrow> (f has_derivative g') F" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 37 | by simp | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 38 | |
| 63558 | 39 | definition has_field_derivative :: "('a::real_normed_field \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 40 | (infix "(has'_field'_derivative)" 50) | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 41 | where "(f has_field_derivative D) F \<longleftrightarrow> (f has_derivative (*) D) F" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 42 | |
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 43 | lemma DERIV_cong: "(f has_field_derivative X) F \<Longrightarrow> X = Y \<Longrightarrow> (f has_field_derivative Y) F" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 44 | by simp | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 45 | |
| 63558 | 46 | definition has_vector_derivative :: "(real \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> real filter \<Rightarrow> bool" | 
| 47 | (infix "has'_vector'_derivative" 50) | |
| 48 | where "(f has_vector_derivative f') net \<longleftrightarrow> (f has_derivative (\<lambda>x. x *\<^sub>R f')) net" | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 49 | |
| 63558 | 50 | lemma has_vector_derivative_eq_rhs: | 
| 51 | "(f has_vector_derivative X) F \<Longrightarrow> X = Y \<Longrightarrow> (f has_vector_derivative Y) F" | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 52 | by simp | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 53 | |
| 57953 | 54 | named_theorems derivative_intros "structural introduction rules for derivatives" | 
| 60758 | 55 | setup \<open> | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 56 | let | 
| 57953 | 57 |     val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs}
 | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 58 | fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 59 | in | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 60 | Global_Theory.add_thms_dynamic | 
| 67149 | 61 | (\<^binding>\<open>derivative_eq_intros\<close>, | 
| 57953 | 62 | fn context => | 
| 69593 | 63 | Named_Theorems.get (Context.proof_of context) \<^named_theorems>\<open>derivative_intros\<close> | 
| 57953 | 64 | |> map_filter eq_rule) | 
| 69216 
1a52baa70aed
clarified ML_Context.expression: it is a closed expression, not a let-declaration -- thus source positions are more accurate (amending d8849cfad60f, 162a4c2e97bc);
 wenzelm parents: 
69111diff
changeset | 65 | end | 
| 60758 | 66 | \<close> | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 67 | |
| 60758 | 68 | text \<open> | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 69 | The following syntax is only used as a legacy syntax. | 
| 60758 | 70 | \<close> | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 71 | abbreviation (input) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 72 |   FDERIV :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow>  ('a \<Rightarrow> 'b) \<Rightarrow> bool"
 | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 73 |   ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
 | 
| 63558 | 74 | where "FDERIV f x :> f' \<equiv> (f has_derivative f') (at x)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 75 | |
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 76 | lemma has_derivative_bounded_linear: "(f has_derivative f') F \<Longrightarrow> bounded_linear 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 | 77 | 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 | 78 | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56289diff
changeset | 79 | lemma has_derivative_linear: "(f has_derivative f') F \<Longrightarrow> linear f'" | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56289diff
changeset | 80 | using bounded_linear.linear[OF has_derivative_bounded_linear] . | 
| 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56289diff
changeset | 81 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 82 | lemma has_derivative_ident[derivative_intros, simp]: "((\<lambda>x. x) has_derivative (\<lambda>x. x)) F" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 83 | by (simp add: has_derivative_def) | 
| 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 | 84 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63299diff
changeset | 85 | lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63299diff
changeset | 86 | by (metis eq_id_iff has_derivative_ident) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63299diff
changeset | 87 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 88 | lemma has_derivative_const[derivative_intros, simp]: "((\<lambda>x. c) has_derivative (\<lambda>x. 0)) F" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 89 | by (simp add: has_derivative_def) | 
| 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 | 90 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; 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 | 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 | 92 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 93 | lemma (in bounded_linear) has_derivative: | 
| 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 | 94 | "(g has_derivative g') F \<Longrightarrow> ((\<lambda>x. f (g x)) has_derivative (\<lambda>x. f (g' x))) F" | 
| 63092 | 95 | unfolding has_derivative_def | 
| 68634 | 96 | by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) | 
| 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 | 97 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 98 | lemmas has_derivative_scaleR_right [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 99 | bounded_linear.has_derivative [OF bounded_linear_scaleR_right] | 
| 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 | 100 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 101 | lemmas has_derivative_scaleR_left [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 102 | bounded_linear.has_derivative [OF bounded_linear_scaleR_left] | 
| 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 | 103 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 104 | lemmas has_derivative_mult_right [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 105 | bounded_linear.has_derivative [OF bounded_linear_mult_right] | 
| 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 | 106 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 107 | lemmas has_derivative_mult_left [derivative_intros] = | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 108 | bounded_linear.has_derivative [OF bounded_linear_mult_left] | 
| 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 | 109 | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 110 | lemmas has_derivative_of_real[derivative_intros, simp] = | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 111 | bounded_linear.has_derivative[OF bounded_linear_of_real] | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 112 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 113 | lemma has_derivative_add[simp, derivative_intros]: | 
| 63558 | 114 | assumes f: "(f has_derivative f') F" | 
| 115 | and g: "(g has_derivative g') 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 | 116 | 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 | 117 | 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 | 118 | 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 | 119 | 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 | 120 | let ?D = "\<lambda>f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" | 
| 61973 | 121 | have "((\<lambda>x. ?D f f' x + ?D g g' x) \<longlongrightarrow> (0 + 0)) 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 | 122 | using f g by (intro tendsto_add) (auto simp: has_derivative_def) | 
| 61973 | 123 | then show "(?D (\<lambda>x. f x + g x) (\<lambda>x. f' x + g' x) \<longlongrightarrow> 0) 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 | 124 | by (simp add: field_simps scaleR_add_right scaleR_diff_right) | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 125 | qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) | 
| 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 | 126 | |
| 64267 | 127 | lemma has_derivative_sum[simp, derivative_intros]: | 
| 63915 | 128 | "(\<And>i. i \<in> I \<Longrightarrow> (f i has_derivative f' i) F) \<Longrightarrow> | 
| 129 | ((\<lambda>x. \<Sum>i\<in>I. f i x) has_derivative (\<lambda>x. \<Sum>i\<in>I. f' i x)) F" | |
| 130 | by (induct I rule: infinite_finite_induct) simp_all | |
| 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 | 131 | |
| 63558 | 132 | lemma has_derivative_minus[simp, derivative_intros]: | 
| 133 | "(f has_derivative f') F \<Longrightarrow> ((\<lambda>x. - f x) has_derivative (\<lambda>x. - f' x)) F" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 134 | using has_derivative_scaleR_right[of f f' F "-1"] by simp | 
| 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 | 135 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 136 | lemma has_derivative_diff[simp, derivative_intros]: | 
| 63558 | 137 | "(f has_derivative f') F \<Longrightarrow> (g has_derivative g') F \<Longrightarrow> | 
| 138 | ((\<lambda>x. f x - g x) has_derivative (\<lambda>x. f' x - g' x)) F" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 139 | by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_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 | 140 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 141 | lemma has_derivative_at_within: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 142 | "(f has_derivative f') (at x within s) \<longleftrightarrow> | 
| 61973 | 143 | (bounded_linear f' \<and> ((\<lambda>y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within s))" | 
| 72219 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 144 | proof (cases "at x within s = bot") | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 145 | case True | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 146 | then show ?thesis | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 147 | by (metis (no_types, lifting) has_derivative_within tendsto_bot) | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 148 | next | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 149 | case False | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 150 | then show ?thesis | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 151 | by (simp add: Lim_ident_at has_derivative_def) | 
| 
0f38c96a0a74
tidying up some theorem statements
 paulson <lp15@cam.ac.uk> parents: 
71837diff
changeset | 152 | qed | 
| 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 | 153 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 154 | lemma has_derivative_iff_norm: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 155 | "(f has_derivative f') (at x within s) \<longleftrightarrow> | 
| 63558 | 156 | bounded_linear f' \<and> ((\<lambda>y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \<longlongrightarrow> 0) (at x within s)" | 
| 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 | 157 | using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 158 | by (simp add: has_derivative_at_within divide_inverse ac_simps) | 
| 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 | 159 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 160 | lemma has_derivative_at: | 
| 63558 | 161 | "(f has_derivative D) (at x) \<longleftrightarrow> | 
| 162 | (bounded_linear D \<and> (\<lambda>h. norm (f (x + h) - f x - D h) / norm h) \<midarrow>0\<rightarrow> 0)" | |
| 72245 | 163 | by (simp add: has_derivative_iff_norm LIM_offset_zero_iff) | 
| 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 | 164 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 165 | lemma field_has_derivative_at: | 
| 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 | 166 | fixes x :: "'a::real_normed_field" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 167 | shows "(f has_derivative (*) D) (at x) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow> D" (is "?lhs = ?rhs") | 
| 68634 | 168 | proof - | 
| 169 | have "?lhs = (\<lambda>h. norm (f (x + h) - f x - D * h) / norm h) \<midarrow>0 \<rightarrow> 0" | |
| 170 | by (simp add: bounded_linear_mult_right has_derivative_at) | |
| 171 | also have "... = (\<lambda>y. norm ((f (x + y) - f x - D * y) / y)) \<midarrow>0\<rightarrow> 0" | |
| 172 | by (simp cong: LIM_cong flip: nonzero_norm_divide) | |
| 173 | also have "... = (\<lambda>y. norm ((f (x + y) - f x) / y - D / y * y)) \<midarrow>0\<rightarrow> 0" | |
| 174 | by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) | |
| 175 | also have "... = ?rhs" | |
| 176 | by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) | |
| 177 | finally show ?thesis . | |
| 178 | qed | |
| 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 | 179 | |
| 70999 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 180 | lemma has_derivative_iff_Ex: | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 181 | "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 182 | bounded_linear f' \<and> (\<exists>e. (\<forall>h. f (x+h) = f x + f' h + e h) \<and> ((\<lambda>h. norm (e h) / norm h) \<longlongrightarrow> 0) (at 0))" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 183 | unfolding has_derivative_at by force | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 184 | |
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 185 | lemma has_derivative_at_within_iff_Ex: | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 186 | assumes "x \<in> S" "open S" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 187 | shows "(f has_derivative f') (at x within S) \<longleftrightarrow> | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 188 | bounded_linear f' \<and> (\<exists>e. (\<forall>h. x+h \<in> S \<longrightarrow> f (x+h) = f x + f' h + e h) \<and> ((\<lambda>h. norm (e h) / norm h) \<longlongrightarrow> 0) (at 0))" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 189 | (is "?lhs = ?rhs") | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 190 | proof safe | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 191 | show "bounded_linear f'" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 192 | if "(f has_derivative f') (at x within S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 193 | using has_derivative_bounded_linear that by blast | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 194 | show "\<exists>e. (\<forall>h. x + h \<in> S \<longrightarrow> f (x + h) = f x + f' h + e h) \<and> (\<lambda>h. norm (e h) / norm h) \<midarrow>0\<rightarrow> 0" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 195 | if "(f has_derivative f') (at x within S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 196 | by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 197 | show "(f has_derivative f') (at x within S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 198 | if "bounded_linear f'" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 199 | and eq [rule_format]: "\<forall>h. x + h \<in> S \<longrightarrow> f (x + h) = f x + f' h + e h" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 200 | and 0: "(\<lambda>h. norm (e (h::'a)::'b) / norm h) \<midarrow>0\<rightarrow> 0" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 201 | for e | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 202 | proof - | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 203 | have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \<in> S" for y | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 204 | using eq [of "y-x"] that by simp | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 205 | have 2: "((\<lambda>y. norm (e (y-x)) / norm (y - x)) \<longlongrightarrow> 0) (at x within S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 206 | by (simp add: "0" assms tendsto_offset_zero_iff) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 207 | have "((\<lambda>y. norm (f y - f x - f' (y - x)) / norm (y - x)) \<longlongrightarrow> 0) (at x within S)" | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 208 | by (simp add: Lim_cong_within 1 2) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 209 | then show ?thesis | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 210 | by (simp add: has_derivative_iff_norm \<open>bounded_linear f'\<close>) | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 211 | qed | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 212 | qed | 
| 
5b753486c075
Inverse function theorem + lemmas
 paulson <lp15@cam.ac.uk> parents: 
70707diff
changeset | 213 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 214 | lemma has_derivativeI: | 
| 63558 | 215 | "bounded_linear f' \<Longrightarrow> | 
| 216 | ((\<lambda>y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \<longlongrightarrow> 0) (at x within s) \<Longrightarrow> | |
| 217 | (f has_derivative f') (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 218 | by (simp add: has_derivative_at_within) | 
| 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 | 219 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 220 | lemma has_derivativeI_sandwich: | 
| 63558 | 221 | assumes e: "0 < e" | 
| 222 | and bounded: "bounded_linear f'" | |
| 223 | and sandwich: "(\<And>y. y \<in> s \<Longrightarrow> y \<noteq> x \<Longrightarrow> dist y x < e \<Longrightarrow> | |
| 224 | norm ((f y - f x) - f' (y - x)) / norm (y - x) \<le> H y)" | |
| 61973 | 225 | and "(H \<longlongrightarrow> 0) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 226 | shows "(f has_derivative f') (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 227 | unfolding has_derivative_iff_norm | 
| 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 | 228 | proof safe | 
| 61973 | 229 | show "((\<lambda>y. norm (f y - f x - f' (y - x)) / norm (y - x)) \<longlongrightarrow> 0) (at x within s)" | 
| 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 | 230 | proof (rule tendsto_sandwich[where f="\<lambda>x. 0"]) | 
| 61973 | 231 | show "(H \<longlongrightarrow> 0) (at x within s)" by fact | 
| 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 | 232 | 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 | 233 | unfolding eventually_at using e sandwich by auto | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 234 | qed (auto simp: le_divide_eq) | 
| 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 | 235 | 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 | 236 | |
| 63558 | 237 | lemma has_derivative_subset: | 
| 238 | "(f has_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_derivative f') (at x within t)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 239 | by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) | 
| 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 | 240 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 241 | lemma has_derivative_within_singleton_iff: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 242 |   "(f has_derivative g) (at x within {x}) \<longleftrightarrow> bounded_linear g"
 | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 243 | by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 244 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 245 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 246 | subsubsection \<open>Limit transformation for derivatives\<close> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 247 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 248 | lemma has_derivative_transform_within: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 249 | assumes "(f has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 250 | and "0 < d" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 251 | and "x \<in> s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 252 | and "\<And>x'. \<lbrakk>x' \<in> s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 253 | shows "(g has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 254 | using assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 255 | unfolding has_derivative_within | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 256 | by (force simp add: intro: Lim_transform_within) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 257 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 258 | lemma has_derivative_transform_within_open: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 259 | assumes "(f has_derivative f') (at x within t)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 260 | and "open s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 261 | and "x \<in> s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 262 | and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 263 | shows "(g has_derivative f') (at x within t)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 264 | using assms unfolding has_derivative_within | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 265 | by (force simp add: intro: Lim_transform_within_open) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 266 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 267 | lemma has_derivative_transform: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 268 | assumes "x \<in> s" "\<And>x. x \<in> s \<Longrightarrow> g x = f x" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 269 | assumes "(f has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 270 | shows "(g has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 271 | using assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 272 | by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 273 | |
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 274 | lemma has_derivative_transform_eventually: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 275 | assumes "(f has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 276 | "(\<forall>\<^sub>F x' in at x within s. f x' = g x')" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 277 | assumes "f x = g x" "x \<in> s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 278 | shows "(g has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 279 | using assms | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 280 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 281 | from assms(2,3) obtain d where "d > 0" "\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x'" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 282 | by (force simp: eventually_at) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 283 | from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 284 | show ?thesis . | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 285 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 286 | |
| 71029 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 287 | lemma has_field_derivative_transform_within: | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 288 | assumes "(f has_field_derivative f') (at a within S)" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 289 | and "0 < d" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 290 | and "a \<in> S" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 291 | and "\<And>x. \<lbrakk>x \<in> S; dist x a < d\<rbrakk> \<Longrightarrow> f x = g x" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 292 | shows "(g has_field_derivative f') (at a within S)" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 293 | using assms unfolding has_field_derivative_def | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 294 | by (metis has_derivative_transform_within) | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 295 | |
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 296 | lemma has_field_derivative_transform_within_open: | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 297 | assumes "(f has_field_derivative f') (at a)" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 298 | and "open S" "a \<in> S" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 299 | and "\<And>x. x \<in> S \<Longrightarrow> f x = g x" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 300 | shows "(g has_field_derivative f') (at a)" | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 301 | using assms unfolding has_field_derivative_def | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 302 | by (metis has_derivative_transform_within_open) | 
| 
934e0044e94b
Moved or deleted some out of place material, also eliminating obsolete naming conventions
 paulson <lp15@cam.ac.uk> parents: 
70999diff
changeset | 303 | |
| 56261 | 304 | |
| 60758 | 305 | subsection \<open>Continuity\<close> | 
| 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 | 306 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 307 | lemma has_derivative_continuous: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 308 | assumes f: "(f has_derivative f') (at x within s)" | 
| 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 | 309 | 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 | 310 | proof - | 
| 63558 | 311 | from f interpret F: bounded_linear f' | 
| 312 | by (rule has_derivative_bounded_linear) | |
| 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 | 313 | note F.tendsto[tendsto_intros] | 
| 61973 | 314 | let ?L = "\<lambda>f. (f \<longlongrightarrow> 0) (at x within s)" | 
| 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 | 315 | have "?L (\<lambda>y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 316 | using f unfolding has_derivative_iff_norm by blast | 
| 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 | 317 | 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 | 318 | 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 | 319 | 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 | 320 | 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 | 321 | 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 | 322 | 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 | 323 | 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 | 324 | 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 | 325 | 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 | 326 | 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 | 327 | 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 | 328 | 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 | 329 | 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 | 330 | |
| 63558 | 331 | |
| 60758 | 332 | subsection \<open>Composition\<close> | 
| 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 | 333 | |
| 63558 | 334 | lemma tendsto_at_iff_tendsto_nhds_within: | 
| 335 | "f x = y \<Longrightarrow> (f \<longlongrightarrow> y) (at x within s) \<longleftrightarrow> (f \<longlongrightarrow> y) (inf (nhds x) (principal s))" | |
| 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 | 336 | unfolding tendsto_def eventually_inf_principal eventually_at_filter | 
| 61810 | 337 | by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) | 
| 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 | 338 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 339 | lemma has_derivative_in_compose: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 340 | assumes f: "(f has_derivative f') (at x within s)" | 
| 63558 | 341 | and g: "(g has_derivative g') (at (f x) within (f`s))" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 342 | shows "((\<lambda>x. g (f x)) has_derivative (\<lambda>x. g' (f' x))) (at x within s)" | 
| 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 | 343 | proof - | 
| 63558 | 344 | from f interpret F: bounded_linear f' | 
| 345 | by (rule has_derivative_bounded_linear) | |
| 346 | from g interpret G: bounded_linear g' | |
| 347 | by (rule has_derivative_bounded_linear) | |
| 348 | from F.bounded obtain kF where kF: "\<And>x. norm (f' x) \<le> norm x * kF" | |
| 349 | by fast | |
| 350 | from G.bounded obtain kG where kG: "\<And>x. norm (g' x) \<le> norm x * kG" | |
| 351 | by fast | |
| 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 | 352 | 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 | 353 | |
| 61973 | 354 | let ?L = "\<lambda>f. (f \<longlongrightarrow> 0) (at x within s)" | 
| 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 | 355 | 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 | 356 | 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 | 357 | let ?gf = "\<lambda>x. g (f x)" and ?gf' = "\<lambda>x. g' (f' x)" | 
| 63040 | 358 | define Nf where "Nf = ?N f f' x" | 
| 359 | define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y | |
| 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 | 360 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; 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 | show ?thesis | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 362 | proof (rule has_derivativeI_sandwich[of 1]) | 
| 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 | 363 | show "bounded_linear (\<lambda>x. g' (f' x))" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 364 | using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) | 
| 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 | 365 | next | 
| 63558 | 366 | fix y :: 'a | 
| 367 | assume neq: "y \<noteq> 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 | 368 | 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 | 369 | 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 | 370 | 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 | 371 | 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 | 372 | 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 | 373 | 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 | 374 | 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 | 375 | 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 | 376 | 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 | 377 | 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 | 378 | 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 | 379 | 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 | 380 | 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 | 381 | by (simp add: neq Nf_def field_simps) | 
| 63558 | 382 | qed (use kG in \<open>simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\<close>) | 
| 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 | 383 | 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 | 384 | 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 | 385 | have [tendsto_intros]: "?L Nf" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 386 | using f unfolding has_derivative_iff_norm Nf_def .. | 
| 61973 | 387 | from f have "(f \<longlongrightarrow> f x) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 388 | by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) | 
| 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 | 389 | 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 | 390 | 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 | 391 | 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 | 392 | |
| 61973 | 393 | have "((?N g g' (f x)) \<longlongrightarrow> 0) (at (f x) within f`s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 394 | using g unfolding has_derivative_iff_norm .. | 
| 61973 | 395 | then have g': "((?N g g' (f x)) \<longlongrightarrow> 0) (inf (nhds (f x)) (principal (f`s)))" | 
| 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 | 396 | 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 | 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 | 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 | 399 | unfolding Ng_def by (rule filterlim_compose[OF g' f']) | 
| 61973 | 400 | show "((\<lambda>y. Nf y * kG + Ng y * (Nf y + kF)) \<longlongrightarrow> 0) (at x within s)" | 
| 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 | 401 | 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 | 402 | 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 | 403 | 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 | 404 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 405 | lemma has_derivative_compose: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 406 | "(f has_derivative f') (at x within s) \<Longrightarrow> (g has_derivative g') (at (f x)) \<Longrightarrow> | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 407 | ((\<lambda>x. g (f x)) has_derivative (\<lambda>x. g' (f' x))) (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 408 | by (blast intro: has_derivative_in_compose has_derivative_subset) | 
| 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 | 409 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 410 | lemma has_derivative_in_compose2: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 411 | assumes "\<And>x. x \<in> t \<Longrightarrow> (g has_derivative g' x) (at x within t)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 412 | assumes "f ` s \<subseteq> t" "x \<in> s" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 413 | assumes "(f has_derivative f') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 414 | shows "((\<lambda>x. g (f x)) has_derivative (\<lambda>y. g' (f x) (f' y))) (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 415 | using assms | 
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
72245diff
changeset | 416 | by (auto intro: has_derivative_subset intro!: has_derivative_in_compose[of f f' x s g]) | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 417 | |
| 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 | 418 | lemma (in bounded_bilinear) FDERIV: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 419 | assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 420 | shows "((\<lambda>x. f x ** g x) has_derivative (\<lambda>h. f x ** g' h + f' h ** g x)) (at x within s)" | 
| 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 | proof - | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 422 | from bounded_linear.bounded [OF has_derivative_bounded_linear [OF 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 | 423 | 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 | 424 | |
| 63558 | 425 | from pos_bounded obtain K | 
| 426 | where K: "0 < K" and norm_prod: "\<And>a b. norm (a ** b) \<le> norm a * norm b * K" | |
| 427 | by fast | |
| 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 | 428 | 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 | 429 | let ?N = "\<lambda>f f' y. norm (?D f f' y) / norm (y - x)" | 
| 63040 | 430 | define Ng where "Ng = ?N g g'" | 
| 431 | define Nf where "Nf = ?N f 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 | 432 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; 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 | 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 | 434 | 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 | 435 | let ?F = "at x within s" | 
| 21164 | 436 | |
| 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 | 437 | show ?thesis | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 438 | proof (rule has_derivativeI_sandwich[of 1]) | 
| 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 | 439 | 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 | 440 | 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 | 441 | bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 442 | has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF 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 | 443 | next | 
| 61973 | 444 | from g have "(g \<longlongrightarrow> g x) ?F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 445 | by (intro continuous_within[THEN iffD1] has_derivative_continuous) | 
| 61973 | 446 | moreover from f g have "(Nf \<longlongrightarrow> 0) ?F" "(Ng \<longlongrightarrow> 0) ?F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 447 | by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) | 
| 61973 | 448 | ultimately have "(?fun2 \<longlongrightarrow> norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?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 | 449 | by (intro tendsto_intros) (simp_all add: LIM_zero_iff) | 
| 61973 | 450 | then show "(?fun2 \<longlongrightarrow> 0) ?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 | 451 | 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 | 452 | next | 
| 63558 | 453 | fix y :: 'd | 
| 454 | assume "y \<noteq> x" | |
| 455 | have "?fun1 y = | |
| 456 | norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - 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 | 457 | 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 | 458 | 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 | 459 | 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 | 460 | 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 | 461 | 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 | 462 | 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 | 463 | 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 | 464 | 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 | 465 | 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 | 466 | 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 | 467 | 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 | 468 | 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 | 469 | 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 | 470 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 471 | lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 472 | lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] | 
| 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 | 473 | |
| 64272 | 474 | lemma has_derivative_prod[simp, derivative_intros]: | 
| 63558 | 475 | fixes f :: "'i \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_field" | 
| 68634 | 476 | shows "(\<And>i. i \<in> I \<Longrightarrow> (f i has_derivative f' i) (at x within S)) \<Longrightarrow> | 
| 477 |     ((\<lambda>x. \<Prod>i\<in>I. f i x) has_derivative (\<lambda>y. \<Sum>i\<in>I. f' i y * (\<Prod>j\<in>I - {i}. f j x))) (at x within S)"
 | |
| 63915 | 478 | proof (induct I rule: infinite_finite_induct) | 
| 479 | case infinite | |
| 480 | then show ?case by simp | |
| 481 | next | |
| 482 | case empty | |
| 483 | then show ?case by simp | |
| 63558 | 484 | next | 
| 63915 | 485 | case (insert i I) | 
| 486 |   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)"
 | |
| 68634 | 487 | have "((\<lambda>x. f i x * (\<Prod>i\<in>I. f i x)) has_derivative ?P) (at x within S)" | 
| 63915 | 488 | using insert by (intro has_derivative_mult) auto | 
| 489 |   also have "?P = (\<lambda>y. \<Sum>i'\<in>insert i I. f' i' y * (\<Prod>j\<in>insert i I - {i'}. f j x))"
 | |
| 490 | using insert(1,2) | |
| 64267 | 491 | by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) | 
| 63915 | 492 | finally show ?case | 
| 493 | using insert by simp | |
| 63558 | 494 | qed | 
| 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 | 495 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 496 | lemma has_derivative_power[simp, derivative_intros]: | 
| 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 | 497 | fixes f :: "'a :: real_normed_vector \<Rightarrow> 'b :: real_normed_field" | 
| 68634 | 498 | assumes f: "(f has_derivative f') (at x within S)" | 
| 499 | shows "((\<lambda>x. f x^n) has_derivative (\<lambda>y. of_nat n * f' y * f x^(n - 1))) (at x within S)" | |
| 64272 | 500 |   using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps)
 | 
| 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 | 501 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 502 | lemma has_derivative_inverse': | 
| 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 | 503 | 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 | 504 | assumes x: "x \<noteq> 0" | 
| 68634 | 505 | shows "(inverse has_derivative (\<lambda>h. - (inverse x * h * inverse x))) (at x within S)" | 
| 506 | (is "(_ has_derivative ?f) _") | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 507 | proof (rule has_derivativeI_sandwich) | 
| 68634 | 508 | show "bounded_linear (\<lambda>h. - (inverse x * h * inverse x))" | 
| 509 | by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) | |
| 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 | 510 | show "0 < norm x" using x by simp | 
| 68634 | 511 | have "(inverse \<longlongrightarrow> inverse x) (at x within S)" | 
| 512 | using tendsto_inverse tendsto_ident_at x by auto | |
| 513 | then show "((\<lambda>y. norm (inverse y - inverse x) * norm (inverse x)) \<longlongrightarrow> 0) (at x within S)" | |
| 514 | by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) | |
| 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 | 515 | next | 
| 63558 | 516 | fix y :: 'a | 
| 517 | assume h: "y \<noteq> x" "dist y x < norm x" | |
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
61976diff
changeset | 518 | then have "y \<noteq> 0" by auto | 
| 68634 | 519 | have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) | 
| 520 | = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / | |
| 521 | norm (y - x)" | |
| 522 | by (simp add: \<open>y \<noteq> 0\<close> inverse_diff_inverse x) | |
| 523 | also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" | |
| 524 | by (simp add: left_diff_distrib norm_minus_commute) | |
| 525 | also have "\<dots> \<le> norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" | |
| 526 | by (simp add: norm_mult) | |
| 527 | also have "\<dots> = norm (inverse y - inverse x) * norm (inverse 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 | 528 | by simp | 
| 68634 | 529 | finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \<le> | 
| 530 | norm (inverse y - inverse x) * norm (inverse 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 | 531 | 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 | 532 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 533 | lemma has_derivative_inverse[simp, derivative_intros]: | 
| 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 | fixes f :: "_ \<Rightarrow> 'a::real_normed_div_algebra" | 
| 63558 | 535 | assumes x: "f x \<noteq> 0" | 
| 68634 | 536 | and f: "(f has_derivative f') (at x within S)" | 
| 63558 | 537 | shows "((\<lambda>x. inverse (f x)) has_derivative (\<lambda>h. - (inverse (f x) * f' h * inverse (f x)))) | 
| 68634 | 538 | (at x within S)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 539 | using has_derivative_compose[OF f has_derivative_inverse', OF 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 | 540 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 541 | lemma has_derivative_divide[simp, derivative_intros]: | 
| 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 | 542 | fixes f :: "_ \<Rightarrow> 'a::real_normed_div_algebra" | 
| 68634 | 543 | assumes f: "(f has_derivative f') (at x within S)" | 
| 544 | and g: "(g has_derivative g') (at x within S)" | |
| 55967 | 545 | assumes x: "g x \<noteq> 0" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 546 | shows "((\<lambda>x. f x / g x) has_derivative | 
| 68634 | 547 | (\<lambda>h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 548 | using has_derivative_mult[OF f has_derivative_inverse[OF x g]] | 
| 56480 
093ea91498e6
field_simps: better support for negation and division, and power
 hoelzl parents: 
56479diff
changeset | 549 | by (simp add: field_simps) | 
| 55967 | 550 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 551 | lemma has_derivative_power_int': | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 552 | fixes x :: "'a::real_normed_field" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 553 | assumes x: "x \<noteq> 0" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 554 | shows "((\<lambda>x. power_int x n) has_derivative (\<lambda>y. y * (of_int n * power_int x (n - 1)))) (at x within S)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 555 | proof (cases n rule: int_cases4) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 556 | case (nonneg n) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 557 | thus ?thesis using x | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 558 | by (cases "n = 0") (auto intro!: derivative_eq_intros simp: field_simps power_int_diff fun_eq_iff | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 559 | simp flip: power_Suc) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 560 | next | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 561 | case (neg n) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 562 | thus ?thesis using x | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 563 | by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 564 | simp flip: power_Suc power_Suc2 power_add) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 565 | qed | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 566 | |
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 567 | lemma has_derivative_power_int[simp, derivative_intros]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 568 | fixes f :: "_ \<Rightarrow> 'a::real_normed_field" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 569 | assumes x: "f x \<noteq> 0" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 570 | and f: "(f has_derivative f') (at x within S)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 571 | shows "((\<lambda>x. power_int (f x) n) has_derivative (\<lambda>h. f' h * (of_int n * power_int (f x) (n - 1)))) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 572 | (at x within S)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 573 | using has_derivative_compose[OF f has_derivative_power_int', OF x] . | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 574 | |
| 63558 | 575 | |
| 576 | text \<open>Conventional form requires mult-AC laws. Types real and complex only.\<close> | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 577 | |
| 63558 | 578 | lemma has_derivative_divide'[derivative_intros]: | 
| 55967 | 579 | fixes f :: "_ \<Rightarrow> 'a::real_normed_field" | 
| 68634 | 580 | assumes f: "(f has_derivative f') (at x within S)" | 
| 581 | and g: "(g has_derivative g') (at x within S)" | |
| 63558 | 582 | and x: "g x \<noteq> 0" | 
| 68634 | 583 | shows "((\<lambda>x. f x / g x) has_derivative (\<lambda>h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" | 
| 55967 | 584 | proof - | 
| 63558 | 585 | have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = | 
| 586 | (f' h * g x - f x * g' h) / (g x * g x)" for h | |
| 587 | by (simp add: field_simps x) | |
| 55967 | 588 | then show ?thesis | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 589 | using has_derivative_divide [OF f g] x | 
| 55967 | 590 | by simp | 
| 591 | qed | |
| 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 | 592 | |
| 63558 | 593 | |
| 60758 | 594 | subsection \<open>Uniqueness\<close> | 
| 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 | 595 | |
| 60758 | 596 | text \<open> | 
| 69593 | 597 | This can not generally shown for \<^const>\<open>has_derivative\<close>, as we need to approach the point from | 
| 63627 | 598 | all directions. There is a proof in \<open>Analysis\<close> for \<open>euclidean_space\<close>. | 
| 60758 | 599 | \<close> | 
| 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 | 600 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 601 | lemma has_derivative_at2: "(f has_derivative f') (at x) \<longleftrightarrow> | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 602 | bounded_linear f' \<and> ((\<lambda>y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \<longlongrightarrow> 0) (at x)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 603 | using has_derivative_within [of f f' x UNIV] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 604 | by simp | 
| 71827 | 605 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 606 | lemma has_derivative_zero_unique: | 
| 63558 | 607 | assumes "((\<lambda>x. 0) has_derivative F) (at x)" | 
| 608 | shows "F = (\<lambda>h. 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 | 609 | 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 | 610 | interpret F: bounded_linear F | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 611 | using assms by (rule has_derivative_bounded_linear) | 
| 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 | 612 | let ?r = "\<lambda>h. norm (F h) / norm h" | 
| 61976 | 613 | have *: "?r \<midarrow>0\<rightarrow> 0" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 614 | using assms unfolding has_derivative_at by simp | 
| 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 | 615 | 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 | 616 | proof | 
| 63558 | 617 | show "F h = 0" for 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 | 618 | proof (rule ccontr) | 
| 63558 | 619 | assume **: "\<not> ?thesis" | 
| 620 | then have h: "h \<noteq> 0" | |
| 621 | by (auto simp add: F.zero) | |
| 622 | with ** have "0 < ?r h" | |
| 623 | by simp | |
| 68634 | 624 | from LIM_D [OF * this] obtain S | 
| 625 | where S: "0 < S" and r: "\<And>x. x \<noteq> 0 \<Longrightarrow> norm x < S \<Longrightarrow> ?r x < ?r h" | |
| 63558 | 626 | by auto | 
| 68634 | 627 | from dense [OF S] obtain t where t: "0 < t \<and> t < S" .. | 
| 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 | 628 | let ?x = "scaleR (t / norm h) h" | 
| 68634 | 629 | have "?x \<noteq> 0" and "norm ?x < S" | 
| 63558 | 630 | using t h by simp_all | 
| 631 | then have "?r ?x < ?r h" | |
| 632 | by (rule r) | |
| 633 | then show False | |
| 634 | using t h by (simp add: F.scaleR) | |
| 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 | 635 | 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 | 636 | 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 | 637 | 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 | 638 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 639 | lemma has_derivative_unique: | 
| 63558 | 640 | assumes "(f has_derivative F) (at x)" | 
| 641 | and "(f has_derivative F') (at x)" | |
| 642 | shows "F = 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 | 643 | proof - | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 644 | have "((\<lambda>x. 0) has_derivative (\<lambda>h. F h - F' h)) (at x)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 645 | using has_derivative_diff [OF assms] by simp | 
| 63558 | 646 | then have "(\<lambda>h. F h - F' h) = (\<lambda>h. 0)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 647 | by (rule has_derivative_zero_unique) | 
| 63558 | 648 | then show "F = 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 | 649 | 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 | 650 | 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 | 651 | |
| 71827 | 652 | lemma has_derivative_Uniq: "\<exists>\<^sub>\<le>\<^sub>1F. (f has_derivative F) (at x)" | 
| 653 | by (simp add: Uniq_def has_derivative_unique) | |
| 654 | ||
| 63558 | 655 | |
| 60758 | 656 | subsection \<open>Differentiability predicate\<close> | 
| 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 | 657 | |
| 63558 | 658 | definition differentiable :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a filter \<Rightarrow> bool"
 | 
| 659 | (infix "differentiable" 50) | |
| 660 | where "f differentiable F \<longleftrightarrow> (\<exists>D. (f has_derivative D) 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 | 661 | |
| 63558 | 662 | lemma differentiable_subset: | 
| 663 | "f differentiable (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f differentiable (at x within t)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 664 | unfolding differentiable_def by (blast intro: has_derivative_subset) | 
| 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 | 665 | |
| 56261 | 666 | lemmas differentiable_within_subset = differentiable_subset | 
| 667 | ||
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 668 | lemma differentiable_ident [simp, derivative_intros]: "(\<lambda>x. x) differentiable F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 669 | unfolding differentiable_def by (blast intro: has_derivative_ident) | 
| 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 | 670 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 671 | lemma differentiable_const [simp, derivative_intros]: "(\<lambda>z. a) differentiable F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 672 | unfolding differentiable_def by (blast intro: has_derivative_const) | 
| 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 | 673 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; 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 | lemma differentiable_in_compose: | 
| 63558 | 675 | "f differentiable (at (g x) within (g`s)) \<Longrightarrow> g differentiable (at x within s) \<Longrightarrow> | 
| 676 | (\<lambda>x. f (g x)) differentiable (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 677 | unfolding differentiable_def by (blast intro: has_derivative_in_compose) | 
| 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 | 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 differentiable_compose: | 
| 63558 | 680 | "f differentiable (at (g x)) \<Longrightarrow> g differentiable (at x within s) \<Longrightarrow> | 
| 681 | (\<lambda>x. f (g x)) differentiable (at x within s)" | |
| 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 | 682 | 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 | 683 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 684 | lemma differentiable_add [simp, derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 685 | "f differentiable F \<Longrightarrow> g differentiable F \<Longrightarrow> (\<lambda>x. f x + g x) differentiable F" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 686 | unfolding differentiable_def by (blast intro: has_derivative_add) | 
| 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 | 687 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 688 | lemma differentiable_sum[simp, derivative_intros]: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 689 | assumes "finite s" "\<forall>a\<in>s. (f a) differentiable net" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 690 | shows "(\<lambda>x. sum (\<lambda>a. f a x) s) differentiable net" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 691 | proof - | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 692 | from bchoice[OF assms(2)[unfolded differentiable_def]] | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 693 | show ?thesis | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 694 | by (auto intro!: has_derivative_sum simp: differentiable_def) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 695 | qed | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 696 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 697 | lemma differentiable_minus [simp, derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 698 | "f differentiable F \<Longrightarrow> (\<lambda>x. - f x) differentiable F" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 699 | unfolding differentiable_def by (blast intro: has_derivative_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 | 700 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 701 | lemma differentiable_diff [simp, derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 702 | "f differentiable F \<Longrightarrow> g differentiable F \<Longrightarrow> (\<lambda>x. f x - g x) differentiable F" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 703 | unfolding differentiable_def by (blast intro: has_derivative_diff) | 
| 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 | 704 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 705 | lemma differentiable_mult [simp, derivative_intros]: | 
| 63558 | 706 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | 
| 707 | shows "f differentiable (at x within s) \<Longrightarrow> g differentiable (at x within s) \<Longrightarrow> | |
| 708 | (\<lambda>x. f x * g x) differentiable (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 709 | unfolding differentiable_def by (blast intro: has_derivative_mult) | 
| 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 | 710 | |
| 73795 | 711 | lemma differentiable_cmult_left_iff [simp]: | 
| 712 | fixes c::"'a::real_normed_field" | |
| 713 | shows "(\<lambda>t. c * q t) differentiable at t \<longleftrightarrow> c = 0 \<or> (\<lambda>t. q t) differentiable at t" (is "?lhs = ?rhs") | |
| 714 | proof | |
| 715 | assume L: ?lhs | |
| 716 |   {assume "c \<noteq> 0"
 | |
| 717 | then have "q differentiable at t" | |
| 718 | using differentiable_mult [OF differentiable_const L, of concl: "1/c"] by auto | |
| 719 | } then show ?rhs | |
| 720 | by auto | |
| 721 | qed auto | |
| 722 | ||
| 723 | lemma differentiable_cmult_right_iff [simp]: | |
| 724 | fixes c::"'a::real_normed_field" | |
| 725 | shows "(\<lambda>t. q t * c) differentiable at t \<longleftrightarrow> c = 0 \<or> (\<lambda>t. q t) differentiable at t" (is "?lhs = ?rhs") | |
| 726 | by (simp add: mult.commute flip: differentiable_cmult_left_iff) | |
| 727 | ||
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 728 | lemma differentiable_inverse [simp, derivative_intros]: | 
| 63558 | 729 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_field" | 
| 730 | shows "f differentiable (at x within s) \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> | |
| 731 | (\<lambda>x. inverse (f x)) differentiable (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 732 | unfolding differentiable_def by (blast intro: has_derivative_inverse) | 
| 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 | 733 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 734 | lemma differentiable_divide [simp, derivative_intros]: | 
| 63558 | 735 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_field" | 
| 736 | shows "f differentiable (at x within s) \<Longrightarrow> g differentiable (at x within s) \<Longrightarrow> | |
| 737 | g x \<noteq> 0 \<Longrightarrow> (\<lambda>x. f x / g x) differentiable (at x within s)" | |
| 63092 | 738 | unfolding divide_inverse by simp | 
| 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 | 739 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 740 | lemma differentiable_power [simp, derivative_intros]: | 
| 63558 | 741 | fixes f g :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_field" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 742 | shows "f differentiable (at x within s) \<Longrightarrow> (\<lambda>x. f x ^ n) differentiable (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 743 | unfolding differentiable_def by (blast intro: has_derivative_power) | 
| 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 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 745 | lemma differentiable_power_int [simp, derivative_intros]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 746 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_field" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 747 | shows "f differentiable (at x within s) \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 748 | (\<lambda>x. power_int (f x) n) differentiable (at x within s)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 749 | unfolding differentiable_def by (blast intro: has_derivative_power_int) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 750 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 751 | lemma differentiable_scaleR [simp, derivative_intros]: | 
| 63558 | 752 | "f differentiable (at x within s) \<Longrightarrow> g differentiable (at x within s) \<Longrightarrow> | 
| 753 | (\<lambda>x. f x *\<^sub>R g x) differentiable (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 754 | unfolding differentiable_def by (blast intro: has_derivative_scaleR) | 
| 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 | 755 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 756 | lemma has_derivative_imp_has_field_derivative: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 757 | "(f has_derivative D) F \<Longrightarrow> (\<And>x. x * D' = D x) \<Longrightarrow> (f has_field_derivative D') F" | 
| 63558 | 758 | unfolding has_field_derivative_def | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 759 | by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 760 | |
| 63558 | 761 | lemma has_field_derivative_imp_has_derivative: | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 762 | "(f has_field_derivative D) F \<Longrightarrow> (f has_derivative (*) D) F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 763 | by (simp add: has_field_derivative_def) | 
| 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 | 764 | |
| 63558 | 765 | lemma DERIV_subset: | 
| 766 | "(f has_field_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> | |
| 767 | (f has_field_derivative f') (at x within t)" | |
| 72445 
2c2de074832e
tidying and removal of legacy name
 paulson <lp15@cam.ac.uk> parents: 
72245diff
changeset | 768 | by (simp add: has_field_derivative_def has_derivative_subset) | 
| 56261 | 769 | |
| 59862 | 770 | lemma has_field_derivative_at_within: | 
| 63558 | 771 | "(f has_field_derivative f') (at x) \<Longrightarrow> (f has_field_derivative f') (at x within s)" | 
| 59862 | 772 | using DERIV_subset by blast | 
| 773 | ||
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 774 | abbreviation (input) | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 775 |   DERIV :: "('a::real_normed_field \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 | 
| 63558 | 776 |     ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
 | 
| 777 | where "DERIV f x :> D \<equiv> (f has_field_derivative D) (at 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 | 778 | |
| 63558 | 779 | abbreviation has_real_derivative :: "(real \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> real filter \<Rightarrow> bool" | 
| 780 | (infix "(has'_real'_derivative)" 50) | |
| 781 | where "(f has_real_derivative D) F \<equiv> (f has_field_derivative D) F" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 782 | |
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 783 | lemma real_differentiable_def: | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 784 | "f differentiable at x within s \<longleftrightarrow> (\<exists>D. (f has_real_derivative D) (at x within s))" | 
| 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 | 785 | proof safe | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 786 | assume "f differentiable at x within s" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 787 | then obtain f' where *: "(f has_derivative f') (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 788 | unfolding differentiable_def by auto | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 789 | then obtain c where "f' = ((*) c)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 790 | by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 791 | with * show "\<exists>D. (f has_real_derivative D) (at x within s)" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 792 | unfolding has_field_derivative_def by auto | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 793 | qed (auto simp: differentiable_def has_field_derivative_def) | 
| 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 | 794 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 795 | lemma real_differentiableE [elim?]: | 
| 63558 | 796 | assumes f: "f differentiable (at x within s)" | 
| 797 | obtains df where "(f has_real_derivative df) (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 798 | using assms by (auto simp: real_differentiable_def) | 
| 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 | 799 | |
| 63079 | 800 | lemma has_field_derivative_iff: | 
| 801 | "(f has_field_derivative D) (at x within S) \<longleftrightarrow> | |
| 802 | ((\<lambda>y. (f y - f x) / (y - x)) \<longlongrightarrow> D) (at x within S)" | |
| 68634 | 803 | proof - | 
| 804 | have "((\<lambda>y. norm (f y - f x - D * (y - x)) / norm (y - x)) \<longlongrightarrow> 0) (at x within S) | |
| 805 | = ((\<lambda>y. (f y - f x) / (y - x) - D) \<longlongrightarrow> 0) (at x within S)" | |
| 806 | apply (subst tendsto_norm_zero_iff[symmetric], rule filterlim_cong) | |
| 807 | apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) | |
| 808 | done | |
| 809 | then show ?thesis | |
| 810 | by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) | |
| 811 | qed | |
| 21164 | 812 | |
| 63079 | 813 | lemma DERIV_def: "DERIV f x :> D \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow> D" | 
| 814 | unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. | |
| 815 | ||
| 73885 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 816 | lemma field_derivative_lim_unique: | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 817 | assumes f: "(f has_field_derivative df) (at z)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 818 | and s: "s \<longlonglongrightarrow> 0" "\<And>n. s n \<noteq> 0" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 819 | and a: "(\<lambda>n. (f (z + s n) - f z) / s n) \<longlonglongrightarrow> a" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 820 | shows "df = a" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 821 | proof (rule ccontr) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 822 | assume "df \<noteq> a" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 823 | obtain q where qpos: "\<And>\<epsilon>. \<epsilon> > 0 \<Longrightarrow> q \<epsilon> > 0" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 824 | and q: "\<And>\<epsilon> y. \<lbrakk>\<epsilon> > 0; y \<noteq> z; dist y z < q \<epsilon>\<rbrakk> \<Longrightarrow> dist ((f y - f z) / (y - z)) df < \<epsilon>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 825 | using f unfolding LIM_def has_field_derivative_iff by metis | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 826 | obtain NA where NA: "\<And>\<epsilon> n. \<lbrakk>\<epsilon> > 0; n \<ge> NA \<epsilon>\<rbrakk> \<Longrightarrow> dist ((f (z + s n) - f z) / s n) a < \<epsilon>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 827 | using a unfolding LIMSEQ_def by metis | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 828 | obtain NB where NB: "\<And>\<epsilon> n. \<lbrakk>\<epsilon> > 0; n \<ge> NB \<epsilon>\<rbrakk> \<Longrightarrow> norm (s n) < \<epsilon>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 829 | using s unfolding LIMSEQ_def by (metis norm_conv_dist) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 830 | have df: "\<And>\<epsilon> n. \<epsilon> > 0 \<Longrightarrow> \<lbrakk>0 < \<epsilon>; norm (s n) < q \<epsilon>\<rbrakk> \<Longrightarrow> dist ((f (z + s n) - f z) / s n) df < \<epsilon>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 831 | using add_cancel_left_right add_diff_cancel_left' q s | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 832 | by (metis add_diff_cancel_right' dist_diff(1)) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 833 | define \<delta> where "\<delta> \<equiv> dist df a / 2" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 834 | with \<open>df \<noteq> a\<close> have "\<delta> > 0" and \<delta>: "\<delta>+\<delta> \<le> dist df a" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 835 | by auto | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 836 | define N where "N \<equiv> max (NA \<delta>) (NB (q \<delta>))" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 837 | then have "norm (s N) < q \<delta>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 838 | by (simp add: NB \<open>\<delta> > 0\<close> qpos) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 839 | then have "dist ((f (z + s N) - f z) / s N) df < \<delta>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 840 | by (simp add: \<open>0 < \<delta>\<close> df) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 841 | moreover have "dist ((f (z + s N) - f z) / s N) a < \<delta>" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 842 | using NA N_def \<open>0 < \<delta>\<close> by force | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 843 | ultimately have "dist df a < dist df a" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 844 | by (smt (verit, ccfv_SIG) \<delta> dist_commute dist_triangle) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 845 | then show False .. | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 846 | qed | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 847 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 848 | lemma mult_commute_abs: "(\<lambda>x. x * c) = (*) c" | 
| 63558 | 849 | for c :: "'a::ab_semigroup_mult" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 850 | by (simp add: fun_eq_iff mult.commute) | 
| 21164 | 851 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 852 | lemma DERIV_compose_FDERIV: | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 853 | fixes f::"real\<Rightarrow>real" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 854 | assumes "DERIV f (g x) :> f'" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 855 | assumes "(g has_derivative g') (at x within s)" | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 856 | shows "((\<lambda>x. f (g x)) has_derivative (\<lambda>x. g' x * f')) (at x within s)" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 857 | using assms has_derivative_compose[of g g' x s f "(*) f'"] | 
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 858 | by (auto simp: has_field_derivative_def ac_simps) | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 859 | |
| 63558 | 860 | |
| 60758 | 861 | subsection \<open>Vector derivative\<close> | 
| 60177 | 862 | |
| 863 | lemma has_field_derivative_iff_has_vector_derivative: | |
| 864 | "(f has_field_derivative y) F \<longleftrightarrow> (f has_vector_derivative y) F" | |
| 865 | unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. | |
| 866 | ||
| 867 | lemma has_field_derivative_subset: | |
| 63558 | 868 | "(f has_field_derivative y) (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> | 
| 869 | (f has_field_derivative y) (at x within t)" | |
| 60177 | 870 | unfolding has_field_derivative_def by (rule has_derivative_subset) | 
| 871 | ||
| 872 | lemma has_vector_derivative_const[simp, derivative_intros]: "((\<lambda>x. c) has_vector_derivative 0) net" | |
| 873 | by (auto simp: has_vector_derivative_def) | |
| 874 | ||
| 875 | lemma has_vector_derivative_id[simp, derivative_intros]: "((\<lambda>x. x) has_vector_derivative 1) net" | |
| 876 | by (auto simp: has_vector_derivative_def) | |
| 877 | ||
| 878 | lemma has_vector_derivative_minus[derivative_intros]: | |
| 879 | "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. - f x) has_vector_derivative (- f')) net" | |
| 880 | by (auto simp: has_vector_derivative_def) | |
| 881 | ||
| 882 | lemma has_vector_derivative_add[derivative_intros]: | |
| 883 | "(f has_vector_derivative f') net \<Longrightarrow> (g has_vector_derivative g') net \<Longrightarrow> | |
| 884 | ((\<lambda>x. f x + g x) has_vector_derivative (f' + g')) net" | |
| 885 | by (auto simp: has_vector_derivative_def scaleR_right_distrib) | |
| 886 | ||
| 64267 | 887 | lemma has_vector_derivative_sum[derivative_intros]: | 
| 60177 | 888 | "(\<And>i. i \<in> I \<Longrightarrow> (f i has_vector_derivative f' i) net) \<Longrightarrow> | 
| 889 | ((\<lambda>x. \<Sum>i\<in>I. f i x) has_vector_derivative (\<Sum>i\<in>I. f' i)) net" | |
| 64267 | 890 | by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) | 
| 60177 | 891 | |
| 892 | lemma has_vector_derivative_diff[derivative_intros]: | |
| 893 | "(f has_vector_derivative f') net \<Longrightarrow> (g has_vector_derivative g') net \<Longrightarrow> | |
| 894 | ((\<lambda>x. f x - g x) has_vector_derivative (f' - g')) net" | |
| 895 | by (auto simp: has_vector_derivative_def scaleR_diff_right) | |
| 896 | ||
| 61204 | 897 | lemma has_vector_derivative_add_const: | 
| 63558 | 898 | "((\<lambda>t. g t + z) has_vector_derivative f') net = ((\<lambda>t. g t) has_vector_derivative f') net" | 
| 899 | apply (intro iffI) | |
| 68634 | 900 | apply (force dest: has_vector_derivative_diff [where g = "\<lambda>t. z", OF _ has_vector_derivative_const]) | 
| 901 | apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) | |
| 63558 | 902 | done | 
| 61204 | 903 | |
| 904 | lemma has_vector_derivative_diff_const: | |
| 63558 | 905 | "((\<lambda>t. g t - z) has_vector_derivative f') net = ((\<lambda>t. g t) has_vector_derivative f') net" | 
| 906 | using has_vector_derivative_add_const [where z = "-z"] | |
| 907 | by simp | |
| 61204 | 908 | |
| 60177 | 909 | lemma (in bounded_linear) has_vector_derivative: | 
| 910 | assumes "(g has_vector_derivative g') F" | |
| 911 | shows "((\<lambda>x. f (g x)) has_vector_derivative f g') F" | |
| 912 | using has_derivative[OF assms[unfolded has_vector_derivative_def]] | |
| 913 | by (simp add: has_vector_derivative_def scaleR) | |
| 914 | ||
| 915 | lemma (in bounded_bilinear) has_vector_derivative: | |
| 916 | assumes "(f has_vector_derivative f') (at x within s)" | |
| 917 | and "(g has_vector_derivative g') (at x within s)" | |
| 918 | shows "((\<lambda>x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" | |
| 919 | using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] | |
| 920 | by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) | |
| 921 | ||
| 922 | lemma has_vector_derivative_scaleR[derivative_intros]: | |
| 923 | "(f has_field_derivative f') (at x within s) \<Longrightarrow> (g has_vector_derivative g') (at x within s) \<Longrightarrow> | |
| 924 | ((\<lambda>x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" | |
| 925 | unfolding has_field_derivative_iff_has_vector_derivative | |
| 926 | by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) | |
| 927 | ||
| 928 | lemma has_vector_derivative_mult[derivative_intros]: | |
| 929 | "(f has_vector_derivative f') (at x within s) \<Longrightarrow> (g has_vector_derivative g') (at x within s) \<Longrightarrow> | |
| 63558 | 930 | ((\<lambda>x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" | 
| 931 | for f g :: "real \<Rightarrow> 'a::real_normed_algebra" | |
| 60177 | 932 | by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) | 
| 933 | ||
| 934 | lemma has_vector_derivative_of_real[derivative_intros]: | |
| 935 | "(f has_field_derivative D) F \<Longrightarrow> ((\<lambda>x. of_real (f x)) has_vector_derivative (of_real D)) F" | |
| 936 | by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) | |
| 63558 | 937 | (simp add: has_field_derivative_iff_has_vector_derivative) | 
| 60177 | 938 | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 939 | lemma has_vector_derivative_real_field: | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 940 | "(f has_field_derivative f') (at (of_real a)) \<Longrightarrow> ((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a within s)" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 941 | using has_derivative_compose[of of_real of_real a _ f "(*) f'"] | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 942 | by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 943 | |
| 63558 | 944 | lemma has_vector_derivative_continuous: | 
| 945 | "(f has_vector_derivative D) (at x within s) \<Longrightarrow> continuous (at x within s) f" | |
| 60177 | 946 | by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) | 
| 947 | ||
| 70613 | 948 | lemma continuous_on_vector_derivative: | 
| 949 | "(\<And>x. x \<in> S \<Longrightarrow> (f has_vector_derivative f' x) (at x within S)) \<Longrightarrow> continuous_on S f" | |
| 950 | by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) | |
| 951 | ||
| 60177 | 952 | lemma has_vector_derivative_mult_right[derivative_intros]: | 
| 63558 | 953 | fixes a :: "'a::real_normed_algebra" | 
| 60177 | 954 | shows "(f has_vector_derivative x) F \<Longrightarrow> ((\<lambda>x. a * f x) has_vector_derivative (a * x)) F" | 
| 955 | by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) | |
| 956 | ||
| 957 | lemma has_vector_derivative_mult_left[derivative_intros]: | |
| 63558 | 958 | fixes a :: "'a::real_normed_algebra" | 
| 60177 | 959 | shows "(f has_vector_derivative x) F \<Longrightarrow> ((\<lambda>x. f x * a) has_vector_derivative (x * a)) F" | 
| 960 | by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) | |
| 961 | ||
| 74007 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 962 | lemma has_vector_derivative_divide[derivative_intros]: | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 963 | fixes a :: "'a::real_normed_field" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 964 | shows "(f has_vector_derivative x) F \<Longrightarrow> ((\<lambda>x. f x / a) has_vector_derivative (x / a)) F" | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 965 | using has_vector_derivative_mult_left [of f x F "inverse a"] | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 966 | by (simp add: field_class.field_divide_inverse) | 
| 
df976eefcba0
A few new lemmas and simplifications
 paulson <lp15@cam.ac.uk> parents: 
73933diff
changeset | 967 | |
| 60177 | 968 | |
| 60758 | 969 | subsection \<open>Derivatives\<close> | 
| 21164 | 970 | |
| 61976 | 971 | lemma DERIV_D: "DERIV f x :> D \<Longrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow> D" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 972 | by (simp add: DERIV_def) | 
| 21164 | 973 | |
| 63079 | 974 | lemma has_field_derivativeD: | 
| 975 | "(f has_field_derivative D) (at x within S) \<Longrightarrow> | |
| 976 | ((\<lambda>y. (f y - f x) / (y - x)) \<longlongrightarrow> D) (at x within S)" | |
| 977 | by (simp add: has_field_derivative_iff) | |
| 978 | ||
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 979 | lemma DERIV_const [simp, derivative_intros]: "((\<lambda>x. k) has_field_derivative 0) F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 980 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 981 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 982 | lemma DERIV_ident [simp, derivative_intros]: "((\<lambda>x. x) has_field_derivative 1) F" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 983 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto | 
| 21164 | 984 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 985 | lemma field_differentiable_add[derivative_intros]: | 
| 63558 | 986 | "(f has_field_derivative f') F \<Longrightarrow> (g has_field_derivative g') F \<Longrightarrow> | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 987 | ((\<lambda>z. f z + g z) has_field_derivative f' + g') F" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 988 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 989 | (auto simp: has_field_derivative_def field_simps mult_commute_abs) | 
| 56261 | 990 | |
| 991 | corollary DERIV_add: | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 992 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> (g has_field_derivative E) (at x within s) \<Longrightarrow> | 
| 63558 | 993 | ((\<lambda>x. f x + g x) has_field_derivative D + E) (at x within s)" | 
| 56261 | 994 | by (rule field_differentiable_add) | 
| 995 | ||
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 996 | lemma field_differentiable_minus[derivative_intros]: | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 997 | "(f has_field_derivative f') F \<Longrightarrow> ((\<lambda>z. - (f z)) has_field_derivative -f') F" | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 998 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) | 
| 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 999 | (auto simp: has_field_derivative_def field_simps mult_commute_abs) | 
| 21164 | 1000 | |
| 63558 | 1001 | corollary DERIV_minus: | 
| 1002 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | |
| 1003 | ((\<lambda>x. - f x) has_field_derivative -D) (at x within s)" | |
| 56261 | 1004 | by (rule field_differentiable_minus) | 
| 21164 | 1005 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1006 | lemma field_differentiable_diff[derivative_intros]: | 
| 63558 | 1007 | "(f has_field_derivative f') F \<Longrightarrow> | 
| 1008 | (g has_field_derivative g') F \<Longrightarrow> ((\<lambda>z. f z - g z) has_field_derivative f' - g') F" | |
| 63092 | 1009 | by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) | 
| 56261 | 1010 | |
| 1011 | corollary DERIV_diff: | |
| 63558 | 1012 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 1013 | (g has_field_derivative E) (at x within s) \<Longrightarrow> | |
| 1014 | ((\<lambda>x. f x - g x) has_field_derivative D - E) (at x within s)" | |
| 56261 | 1015 | by (rule field_differentiable_diff) | 
| 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 | 1016 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1017 | lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \<Longrightarrow> continuous (at x within s) f" | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1018 | by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp | 
| 21164 | 1019 | |
| 56261 | 1020 | corollary DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x" | 
| 1021 | by (rule DERIV_continuous) | |
| 1022 | ||
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1023 | lemma DERIV_atLeastAtMost_imp_continuous_on: | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1024 | assumes "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> \<exists>y. DERIV f x :> y" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1025 |   shows "continuous_on {a..b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1026 | by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1027 | |
| 56261 | 1028 | lemma DERIV_continuous_on: | 
| 63299 | 1029 | "(\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative (D x)) (at x within s)) \<Longrightarrow> continuous_on s f" | 
| 1030 | unfolding continuous_on_eq_continuous_within | |
| 63558 | 1031 | by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) | 
| 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 | 1032 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1033 | lemma DERIV_mult': | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1034 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> (g has_field_derivative E) (at x within s) \<Longrightarrow> | 
| 63558 | 1035 | ((\<lambda>x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1036 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1037 | (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) | 
| 21164 | 1038 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1039 | lemma DERIV_mult[derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1040 | "(f has_field_derivative Da) (at x within s) \<Longrightarrow> (g has_field_derivative Db) (at x within s) \<Longrightarrow> | 
| 63558 | 1041 | ((\<lambda>x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1042 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1043 | (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) | 
| 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 | 1044 | |
| 60758 | 1045 | text \<open>Derivative of linear multiplication\<close> | 
| 21164 | 1046 | |
| 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 | 1047 | lemma DERIV_cmult: | 
| 63558 | 1048 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 1049 | ((\<lambda>x. c * f x) has_field_derivative c * D) (at x within s)" | |
| 1050 | by (drule DERIV_mult' [OF DERIV_const]) simp | |
| 21164 | 1051 | |
| 55967 | 1052 | lemma DERIV_cmult_right: | 
| 63558 | 1053 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 1054 | ((\<lambda>x. f x * c) has_field_derivative D * c) (at x within s)" | |
| 1055 | using DERIV_cmult by (auto simp add: ac_simps) | |
| 55967 | 1056 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 1057 | lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" | 
| 63558 | 1058 | using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp | 
| 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 | 1059 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1060 | lemma DERIV_cdivide: | 
| 63558 | 1061 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 1062 | ((\<lambda>x. f x / c) has_field_derivative D / c) (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1063 | using DERIV_cmult_right[of f D x s "1 / c"] by simp | 
| 21164 | 1064 | |
| 63558 | 1065 | lemma DERIV_unique: "DERIV f x :> D \<Longrightarrow> DERIV f x :> E \<Longrightarrow> D = E" | 
| 1066 | unfolding DERIV_def by (rule LIM_unique) | |
| 21164 | 1067 | |
| 71827 | 1068 | lemma DERIV_Uniq: "\<exists>\<^sub>\<le>\<^sub>1D. DERIV f x :> D" | 
| 1069 | by (simp add: DERIV_unique Uniq_def) | |
| 1070 | ||
| 64267 | 1071 | lemma DERIV_sum[derivative_intros]: | 
| 63558 | 1072 | "(\<And> n. n \<in> S \<Longrightarrow> ((\<lambda>x. f x n) has_field_derivative (f' x n)) F) \<Longrightarrow> | 
| 64267 | 1073 | ((\<lambda>x. sum (f x) S) has_field_derivative sum (f' x) S) F" | 
| 1074 | by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) | |
| 1075 | (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) | |
| 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 | 1076 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1077 | lemma DERIV_inverse'[derivative_intros]: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1078 | assumes "(f has_field_derivative D) (at x within s)" | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1079 | and "f x \<noteq> 0" | 
| 63558 | 1080 | shows "((\<lambda>x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) | 
| 1081 | (at x within s)" | |
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1082 | proof - | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 1083 | have "(f has_derivative (\<lambda>x. x * D)) = (f has_derivative (*) D)" | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1084 | by (rule arg_cong [of "\<lambda>x. x * D"]) (simp add: fun_eq_iff) | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1085 | with assms have "(f has_derivative (\<lambda>x. x * D)) (at x within s)" | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1086 | by (auto dest!: has_field_derivative_imp_has_derivative) | 
| 60758 | 1087 | then show ?thesis using \<open>f x \<noteq> 0\<close> | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1088 | by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) | 
| 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 1089 | qed | 
| 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 | 1090 | |
| 61799 | 1091 | text \<open>Power of \<open>-1\<close>\<close> | 
| 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 | 1092 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1093 | lemma DERIV_inverse: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1094 | "x \<noteq> 0 \<Longrightarrow> ((\<lambda>x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" | 
| 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 | 1095 | 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 | 1096 | |
| 60758 | 1097 | text \<open>Derivative of inverse\<close> | 
| 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 | 1098 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1099 | lemma DERIV_inverse_fun: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1100 | "(f has_field_derivative d) (at x within s) \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> | 
| 63558 | 1101 | ((\<lambda>x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1102 | by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) | 
| 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 | 1103 | |
| 60758 | 1104 | text \<open>Derivative of quotient\<close> | 
| 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 | 1105 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1106 | lemma DERIV_divide[derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1107 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 63558 | 1108 | (g has_field_derivative E) (at x within s) \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> | 
| 1109 | ((\<lambda>x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1110 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) | 
| 56480 
093ea91498e6
field_simps: better support for negation and division, and power
 hoelzl parents: 
56479diff
changeset | 1111 | (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) | 
| 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 | 1112 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1113 | lemma DERIV_quotient: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1114 | "(f has_field_derivative d) (at x within s) \<Longrightarrow> | 
| 63558 | 1115 | (g has_field_derivative e) (at x within s)\<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> | 
| 1116 | ((\<lambda>y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 1117 | by (drule (2) DERIV_divide) (simp add: mult.commute) | 
| 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 | 1118 | |
| 
400ec5ae7f8f
move FrechetDeriv from the Library to HOL/Deriv; base DERIV on FDERIV and both derivatives allow a restricted support set; FDERIV is now an abbreviation of has_derivative
 hoelzl parents: 
51641diff
changeset | 1119 | lemma DERIV_power_Suc: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1120 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 63558 | 1121 | ((\<lambda>x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1122 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1123 | (auto simp: has_field_derivative_def) | 
| 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 | 1124 | |
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1125 | lemma DERIV_power[derivative_intros]: | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1126 | "(f has_field_derivative D) (at x within s) \<Longrightarrow> | 
| 63558 | 1127 | ((\<lambda>x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1128 | by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1129 | (auto simp: has_field_derivative_def) | 
| 31880 | 1130 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1131 | lemma DERIV_pow: "((\<lambda>x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61552diff
changeset | 1132 | using DERIV_power [OF DERIV_ident] by simp | 
| 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 | 1133 | |
| 71837 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1134 | lemma DERIV_power_int [derivative_intros]: | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1135 | assumes [derivative_intros]: "(f has_field_derivative d) (at x within s)" and [simp]: "f x \<noteq> 0" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1136 | shows "((\<lambda>x. power_int (f x) n) has_field_derivative | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1137 | (of_int n * power_int (f x) (n - 1) * d)) (at x within s)" | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1138 | proof (cases n rule: int_cases4) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1139 | case (nonneg n) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1140 | thus ?thesis | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1141 | by (cases "n = 0") | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1142 | (auto intro!: derivative_eq_intros simp: field_simps power_int_diff | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1143 | simp flip: power_Suc power_Suc2 power_add) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1144 | next | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1145 | case (neg n) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1146 | thus ?thesis | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1147 | by (auto intro!: derivative_eq_intros simp: field_simps power_int_diff power_int_minus | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1148 | simp flip: power_Suc power_Suc2 power_add) | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1149 | qed | 
| 
dca11678c495
new constant power_int in HOL
 Manuel Eberl <eberlm@in.tum.de> parents: 
71827diff
changeset | 1150 | |
| 63558 | 1151 | lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \<Longrightarrow> DERIV g (f x) :> E \<Longrightarrow> | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1152 | ((\<lambda>x. g (f x)) has_field_derivative E * D) (at x within s)" | 
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 1153 | using has_derivative_compose[of f "(*) D" x s g "(*) E"] | 
| 63170 | 1154 | by (simp only: has_field_derivative_def mult_commute_abs ac_simps) | 
| 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 | 1155 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1156 | corollary DERIV_chain2: "DERIV f (g x) :> Da \<Longrightarrow> (g has_field_derivative Db) (at x within s) \<Longrightarrow> | 
| 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1157 | ((\<lambda>x. f (g x)) has_field_derivative Da * Db) (at x within s)" | 
| 55967 | 1158 | by (rule DERIV_chain') | 
| 1159 | ||
| 60758 | 1160 | text \<open>Standard version\<close> | 
| 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 | 1161 | |
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1162 | lemma DERIV_chain: | 
| 63558 | 1163 | "DERIV f (g x) :> Da \<Longrightarrow> (g has_field_derivative Db) (at x within s) \<Longrightarrow> | 
| 1164 | (f \<circ> g has_field_derivative Da * Db) (at x within s)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 1165 | by (drule (1) DERIV_chain', simp add: o_def mult.commute) | 
| 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 | 1166 | |
| 63558 | 1167 | lemma DERIV_image_chain: | 
| 1168 | "(f has_field_derivative Da) (at (g x) within (g ` s)) \<Longrightarrow> | |
| 1169 | (g has_field_derivative Db) (at x within s) \<Longrightarrow> | |
| 1170 | (f \<circ> g has_field_derivative Da * Db) (at x within s)" | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
69022diff
changeset | 1171 | using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1172 | by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) | 
| 55967 | 1173 | |
| 1174 | (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) | |
| 1175 | lemma DERIV_chain_s: | |
| 1176 | assumes "(\<And>x. x \<in> s \<Longrightarrow> DERIV g x :> g'(x))" | |
| 63558 | 1177 | and "DERIV f x :> f'" | 
| 1178 | and "f x \<in> s" | |
| 1179 | shows "DERIV (\<lambda>x. g(f x)) x :> f' * g'(f x)" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57418diff
changeset | 1180 | by (metis (full_types) DERIV_chain' mult.commute assms) | 
| 55967 | 1181 | |
| 1182 | lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) | |
| 1183 | assumes "(\<And>x. DERIV g x :> g'(x))" | |
| 63558 | 1184 | and "DERIV f x :> f'" | 
| 1185 | shows "DERIV (\<lambda>x. g(f x)) x :> f' * g'(f x)" | |
| 55967 | 1186 | by (metis UNIV_I DERIV_chain_s [of UNIV] assms) | 
| 1187 | ||
| 63558 | 1188 | text \<open>Alternative definition for differentiability\<close> | 
| 21164 | 1189 | |
| 1190 | lemma DERIV_LIM_iff: | |
| 63558 | 1191 |   fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a"
 | 
| 68634 | 1192 | shows "((\<lambda>h. (f (a + h) - f a) / h) \<midarrow>0\<rightarrow> D) = ((\<lambda>x. (f x - f a) / (x - a)) \<midarrow>a\<rightarrow> D)" (is "?lhs = ?rhs") | 
| 1193 | proof | |
| 1194 | assume ?lhs | |
| 1195 | then have "(\<lambda>x. (f (a + (x + - a)) - f a) / (x + - a)) \<midarrow>0 - - a\<rightarrow> D" | |
| 1196 | by (rule LIM_offset) | |
| 1197 | then show ?rhs | |
| 1198 | by simp | |
| 1199 | next | |
| 1200 | assume ?rhs | |
| 1201 | then have "(\<lambda>x. (f (x+a) - f a) / ((x+a) - a)) \<midarrow>a-a\<rightarrow> D" | |
| 1202 | by (rule LIM_offset) | |
| 1203 | then show ?lhs | |
| 1204 | by (simp add: add.commute) | |
| 1205 | qed | |
| 63079 | 1206 | |
| 1207 | lemma has_field_derivative_cong_ev: | |
| 1208 | assumes "x = y" | |
| 68635 | 1209 | and *: "eventually (\<lambda>x. x \<in> S \<longrightarrow> f x = g x) (nhds x)" | 
| 1210 | and "u = v" "S = t" "x \<in> S" | |
| 1211 | shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" | |
| 68634 | 1212 | unfolding has_field_derivative_iff | 
| 63079 | 1213 | proof (rule filterlim_cong) | 
| 63558 | 1214 | from assms have "f y = g y" | 
| 1215 | by (auto simp: eventually_nhds) | |
| 68635 | 1216 | with * show "\<forall>\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" | 
| 63079 | 1217 | unfolding eventually_at_filter | 
| 1218 | by eventually_elim (auto simp: assms \<open>f y = g y\<close>) | |
| 1219 | qed (simp_all add: assms) | |
| 21164 | 1220 | |
| 67706 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67443diff
changeset | 1221 | lemma has_field_derivative_cong_eventually: | 
| 68635 | 1222 | assumes "eventually (\<lambda>x. f x = g x) (at x within S)" "f x = g x" | 
| 1223 | shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" | |
| 68634 | 1224 | unfolding has_field_derivative_iff | 
| 68635 | 1225 | proof (rule tendsto_cong) | 
| 1226 | show "\<forall>\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" | |
| 1227 | using assms by (auto elim: eventually_mono) | |
| 1228 | qed | |
| 67706 
4ddc49205f5d
Unified the order of zeros and poles; improved reasoning around non-essential singularites
 Wenda Li <wl302@cam.ac.uk> parents: 
67443diff
changeset | 1229 | |
| 63558 | 1230 | lemma DERIV_cong_ev: | 
| 1231 | "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow> | |
| 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 | 1232 | DERIV f x :> u \<longleftrightarrow> DERIV g y :> v" | 
| 63079 | 1233 | by (rule has_field_derivative_cong_ev) simp_all | 
| 21164 | 1234 | |
| 73885 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1235 | lemma DERIV_mirror: "(DERIV f (- x) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (- x)) x :> - y)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1236 | for f :: "real \<Rightarrow> real" and x y :: real | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1237 | by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1238 | tendsto_minus_cancel_left field_simps conj_commute) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1239 | |
| 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 | 1240 | lemma DERIV_shift: | 
| 63079 | 1241 | "(f has_field_derivative y) (at (x + z)) = ((\<lambda>x. f (x + z)) has_field_derivative y) (at x)" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1242 | by (simp add: DERIV_def field_simps) | 
| 21164 | 1243 | |
| 73885 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1244 | lemma DERIV_at_within_shift_lemma: | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1245 | assumes "(f has_field_derivative y) (at (z+x) within (+) z ` S)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1246 | shows "(f \<circ> (+)z has_field_derivative y) (at x within S)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1247 | proof - | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1248 | have "((+)z has_field_derivative 1) (at x within S)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1249 | by (rule derivative_eq_intros | simp)+ | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1250 | with assms DERIV_image_chain show ?thesis | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1251 | by (metis mult.right_neutral) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1252 | qed | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1253 | |
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1254 | lemma DERIV_at_within_shift: | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1255 | "(f has_field_derivative y) (at (z+x) within (+) z ` S) \<longleftrightarrow> | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1256 | ((\<lambda>x. f (z+x)) has_field_derivative y) (at x within S)" (is "?lhs = ?rhs") | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1257 | proof | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1258 | assume ?lhs then show ?rhs | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1259 | using DERIV_at_within_shift_lemma unfolding o_def by blast | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1260 | next | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1261 | have [simp]: "(\<lambda>x. x - z) ` (+) z ` S = S" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1262 | by force | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1263 | assume R: ?rhs | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1264 | have "(f \<circ> (+) z \<circ> (+) (- z) has_field_derivative y) (at (z + x) within (+) z ` S)" | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1265 | by (rule DERIV_at_within_shift_lemma) (use R in \<open>simp add: o_def\<close>) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1266 | then show ?lhs | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1267 | by (simp add: o_def) | 
| 
26171a89466a
A few useful lemmas about derivatives, colinearity and other topics
 paulson <lp15@cam.ac.uk> parents: 
73795diff
changeset | 1268 | qed | 
| 21164 | 1269 | |
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1270 | lemma floor_has_real_derivative: | 
| 63558 | 1271 |   fixes f :: "real \<Rightarrow> 'a::{floor_ceiling,order_topology}"
 | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1272 | assumes "isCont f x" | 
| 63558 | 1273 | and "f x \<notin> \<int>" | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1274 | shows "((\<lambda>x. floor (f x)) has_real_derivative 0) (at x)" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1275 | proof (subst DERIV_cong_ev[OF refl _ refl]) | 
| 63558 | 1276 | show "((\<lambda>_. floor (f x)) has_real_derivative 0) (at x)" | 
| 1277 | by simp | |
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1278 | have "\<forall>\<^sub>F y in at x. \<lfloor>f y\<rfloor> = \<lfloor>f x\<rfloor>" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1279 | by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1280 | then show "\<forall>\<^sub>F y in nhds x. real_of_int \<lfloor>f y\<rfloor> = real_of_int \<lfloor>f x\<rfloor>" | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1281 | unfolding eventually_at_filter | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1282 | by eventually_elim auto | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1283 | qed | 
| 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1284 | |
| 67685 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 1285 | lemmas has_derivative_floor[derivative_intros] = | 
| 
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
 immler parents: 
67443diff
changeset | 1286 | floor_has_real_derivative[THEN DERIV_compose_FDERIV] | 
| 63263 
c6c95d64607a
approximation, derivative, and continuity of floor and ceiling
 immler parents: 
63170diff
changeset | 1287 | |
| 70707 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1288 | lemma continuous_floor: | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1289 | fixes x::real | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1290 | shows "x \<notin> \<int> \<Longrightarrow> continuous (at x) (real_of_int \<circ> floor)" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1291 | using floor_has_real_derivative [where f=id] | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1292 | by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1293 | |
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1294 | lemma continuous_frac: | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1295 | fixes x::real | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1296 | assumes "x \<notin> \<int>" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1297 | shows "continuous (at x) frac" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1298 | proof - | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1299 | have "isCont (\<lambda>x. real_of_int \<lfloor>x\<rfloor>) x" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1300 | using continuous_floor [OF assms] by (simp add: o_def) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1301 | then have *: "continuous (at x) (\<lambda>x. x - real_of_int \<lfloor>x\<rfloor>)" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1302 | by (intro continuous_intros) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1303 | moreover have "\<forall>\<^sub>F x in nhds x. frac x = x - real_of_int \<lfloor>x\<rfloor>" | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1304 | by (simp add: frac_def) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1305 | ultimately show ?thesis | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1306 | by (simp add: LIM_imp_LIM frac_def isCont_def) | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1307 | qed | 
| 
125705f5965f
A little-known material, and some tidying up
 paulson <lp15@cam.ac.uk> parents: 
70615diff
changeset | 1308 | |
| 60758 | 1309 | text \<open>Caratheodory formulation of derivative at a point\<close> | 
| 21164 | 1310 | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
68638diff
changeset | 1311 | 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 | 1312 | "(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)" | 
| 63558 | 1313 | (is "?lhs = ?rhs") | 
| 21164 | 1314 | proof | 
| 63558 | 1315 | assume ?lhs | 
| 1316 | show "\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> isCont g x \<and> g x = l" | |
| 21164 | 1317 | proof (intro exI conjI) | 
| 63558 | 1318 | let ?g = "(\<lambda>z. if z = x then l else (f z - f x) / (z-x))" | 
| 1319 | show "\<forall>z. f z - f x = ?g z * (z - x)" | |
| 1320 | by simp | |
| 1321 | show "isCont ?g x" | |
| 1322 | using \<open>?lhs\<close> by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) | |
| 1323 | show "?g x = l" | |
| 1324 | by simp | |
| 21164 | 1325 | qed | 
| 1326 | next | |
| 63558 | 1327 | assume ?rhs | 
| 1328 | then show ?lhs | |
| 1329 | by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) | |
| 21164 | 1330 | qed | 
| 1331 | ||
| 1332 | ||
| 60758 | 1333 | subsection \<open>Local extrema\<close> | 
| 29975 | 1334 | |
| 69593 | 1335 | text \<open>If \<^term>\<open>0 < f' x\<close> then \<^term>\<open>x\<close> is Locally Strictly Increasing At The Right.\<close> | 
| 21164 | 1336 | |
| 63079 | 1337 | lemma has_real_derivative_pos_inc_right: | 
| 63558 | 1338 | fixes f :: "real \<Rightarrow> real" | 
| 63079 | 1339 | assumes der: "(f has_real_derivative l) (at x within S)" | 
| 63558 | 1340 | and l: "0 < l" | 
| 63079 | 1341 | shows "\<exists>d > 0. \<forall>h > 0. x + h \<in> S \<longrightarrow> h < d \<longrightarrow> f x < f (x + h)" | 
| 1342 | using assms | |
| 1343 | proof - | |
| 1344 | from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] | |
| 63558 | 1345 | obtain s where s: "0 < s" | 
| 1346 | and all: "\<And>xa. xa\<in>S \<Longrightarrow> xa \<noteq> x \<and> dist xa x < s \<longrightarrow> \<bar>(f xa - f x) / (xa - x) - l\<bar> < l" | |
| 63079 | 1347 | by (auto simp: dist_real_def) | 
| 1348 | then show ?thesis | |
| 1349 | proof (intro exI conjI strip) | |
| 63558 | 1350 | show "0 < s" by (rule s) | 
| 1351 | next | |
| 1352 | fix h :: real | |
| 63079 | 1353 | assume "0 < h" "h < s" "x + h \<in> S" | 
| 1354 | with all [of "x + h"] show "f x < f (x+h)" | |
| 1355 | proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) | |
| 63558 | 1356 | assume "\<not> (f (x + h) - f x) / h < l" and h: "0 < h" | 
| 1357 | with l have "0 < (f (x + h) - f x) / h" | |
| 1358 | by arith | |
| 1359 | then show "f x < f (x + h)" | |
| 63079 | 1360 | by (simp add: pos_less_divide_eq h) | 
| 1361 | qed | |
| 1362 | qed | |
| 1363 | qed | |
| 1364 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1365 | lemma DERIV_pos_inc_right: | 
| 63558 | 1366 | fixes f :: "real \<Rightarrow> real" | 
| 21164 | 1367 | assumes der: "DERIV f x :> l" | 
| 63558 | 1368 | and l: "0 < l" | 
| 1369 | shows "\<exists>d > 0. \<forall>h > 0. h < d \<longrightarrow> f x < f (x + h)" | |
| 63079 | 1370 | using has_real_derivative_pos_inc_right[OF assms] | 
| 1371 | by auto | |
| 1372 | ||
| 1373 | lemma has_real_derivative_neg_dec_left: | |
| 63558 | 1374 | fixes f :: "real \<Rightarrow> real" | 
| 63079 | 1375 | assumes der: "(f has_real_derivative l) (at x within S)" | 
| 63558 | 1376 | and "l < 0" | 
| 63079 | 1377 | shows "\<exists>d > 0. \<forall>h > 0. x - h \<in> S \<longrightarrow> h < d \<longrightarrow> f x < f (x - h)" | 
| 21164 | 1378 | proof - | 
| 63558 | 1379 | from \<open>l < 0\<close> have l: "- l > 0" | 
| 1380 | by simp | |
| 63079 | 1381 | from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] | 
| 63558 | 1382 | obtain s where s: "0 < s" | 
| 1383 | and all: "\<And>xa. xa\<in>S \<Longrightarrow> xa \<noteq> x \<and> dist xa x < s \<longrightarrow> \<bar>(f xa - f x) / (xa - x) - l\<bar> < - l" | |
| 63079 | 1384 | by (auto simp: dist_real_def) | 
| 63558 | 1385 | then show ?thesis | 
| 21164 | 1386 | proof (intro exI conjI strip) | 
| 63558 | 1387 | show "0 < s" by (rule s) | 
| 1388 | next | |
| 1389 | fix h :: real | |
| 63079 | 1390 | assume "0 < h" "h < s" "x - h \<in> S" | 
| 1391 | with all [of "x - h"] show "f x < f (x-h)" | |
| 63648 | 1392 | proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) | 
| 63558 | 1393 | assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" | 
| 1394 | with l have "0 < (f (x-h) - f x) / h" | |
| 1395 | by arith | |
| 1396 | then show "f x < f (x - h)" | |
| 63079 | 1397 | by (simp add: pos_less_divide_eq h) | 
| 21164 | 1398 | qed | 
| 1399 | qed | |
| 1400 | qed | |
| 1401 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1402 | lemma DERIV_neg_dec_left: | 
| 63558 | 1403 | fixes f :: "real \<Rightarrow> real" | 
| 21164 | 1404 | assumes der: "DERIV f x :> l" | 
| 63558 | 1405 | and l: "l < 0" | 
| 1406 | shows "\<exists>d > 0. \<forall>h > 0. h < d \<longrightarrow> f x < f (x - h)" | |
| 63079 | 1407 | using has_real_derivative_neg_dec_left[OF assms] | 
| 1408 | by auto | |
| 1409 | ||
| 1410 | lemma has_real_derivative_pos_inc_left: | |
| 63558 | 1411 | fixes f :: "real \<Rightarrow> real" | 
| 1412 | shows "(f has_real_derivative l) (at x within S) \<Longrightarrow> 0 < l \<Longrightarrow> | |
| 1413 | \<exists>d>0. \<forall>h>0. x - h \<in> S \<longrightarrow> h < d \<longrightarrow> f (x - h) < f x" | |
| 1414 | by (rule has_real_derivative_neg_dec_left [of "\<lambda>x. - f x" "-l" x S, simplified]) | |
| 63079 | 1415 | (auto simp add: DERIV_minus) | 
| 21164 | 1416 | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1417 | lemma DERIV_pos_inc_left: | 
| 63558 | 1418 | fixes f :: "real \<Rightarrow> real" | 
| 1419 | shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d \<longrightarrow> f (x - h) < f x" | |
| 63079 | 1420 | using has_real_derivative_pos_inc_left | 
| 1421 | by blast | |
| 1422 | ||
| 1423 | lemma has_real_derivative_neg_dec_right: | |
| 63558 | 1424 | fixes f :: "real \<Rightarrow> real" | 
| 1425 | shows "(f has_real_derivative l) (at x within S) \<Longrightarrow> l < 0 \<Longrightarrow> | |
| 1426 | \<exists>d > 0. \<forall>h > 0. x + h \<in> S \<longrightarrow> h < d \<longrightarrow> f x > f (x + h)" | |
| 1427 | by (rule has_real_derivative_pos_inc_right [of "\<lambda>x. - f x" "-l" x S, simplified]) | |
| 63079 | 1428 | (auto simp add: DERIV_minus) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1429 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1430 | lemma DERIV_neg_dec_right: | 
| 63558 | 1431 | fixes f :: "real \<Rightarrow> real" | 
| 1432 | shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d \<longrightarrow> f x > f (x + h)" | |
| 63079 | 1433 | using has_real_derivative_neg_dec_right by blast | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1434 | |
| 21164 | 1435 | lemma DERIV_local_max: | 
| 63558 | 1436 | fixes f :: "real \<Rightarrow> real" | 
| 21164 | 1437 | assumes der: "DERIV f x :> l" | 
| 63558 | 1438 | and d: "0 < d" | 
| 1439 | and le: "\<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> f y \<le> f x" | |
| 21164 | 1440 | shows "l = 0" | 
| 1441 | proof (cases rule: linorder_cases [of l 0]) | |
| 63558 | 1442 | case equal | 
| 1443 | then show ?thesis . | |
| 21164 | 1444 | next | 
| 1445 | case less | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1446 | from DERIV_neg_dec_left [OF der less] | 
| 63558 | 1447 | obtain d' where d': "0 < d'" and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x - h)" | 
| 1448 | by blast | |
| 1449 | obtain e where "0 < e \<and> e < d \<and> e < d'" | |
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
67707diff
changeset | 1450 | using field_lbound_gt_zero [OF d d'] .. | 
| 63558 | 1451 | with lt le [THEN spec [where x="x - e"]] show ?thesis | 
| 1452 | by (auto simp add: abs_if) | |
| 21164 | 1453 | next | 
| 1454 | case greater | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1455 | from DERIV_pos_inc_right [OF der greater] | 
| 63558 | 1456 | obtain d' where d': "0 < d'" and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" | 
| 1457 | by blast | |
| 1458 | obtain e where "0 < e \<and> e < d \<and> e < d'" | |
| 68527 
2f4e2aab190a
Generalising and renaming some basic results
 paulson <lp15@cam.ac.uk> parents: 
67707diff
changeset | 1459 | using field_lbound_gt_zero [OF d d'] .. | 
| 63558 | 1460 | with lt le [THEN spec [where x="x + e"]] show ?thesis | 
| 1461 | by (auto simp add: abs_if) | |
| 21164 | 1462 | qed | 
| 1463 | ||
| 63558 | 1464 | text \<open>Similar theorem for a local minimum\<close> | 
| 21164 | 1465 | lemma DERIV_local_min: | 
| 63558 | 1466 | fixes f :: "real \<Rightarrow> real" | 
| 1467 | shows "DERIV f x :> l \<Longrightarrow> 0 < d \<Longrightarrow> \<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> f x \<le> f y \<Longrightarrow> l = 0" | |
| 1468 | by (drule DERIV_minus [THEN DERIV_local_max]) auto | |
| 21164 | 1469 | |
| 1470 | ||
| 60758 | 1471 | text\<open>In particular, if a function is locally flat\<close> | 
| 21164 | 1472 | lemma DERIV_local_const: | 
| 63558 | 1473 | fixes f :: "real \<Rightarrow> real" | 
| 1474 | shows "DERIV f x :> l \<Longrightarrow> 0 < d \<Longrightarrow> \<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> f x = f y \<Longrightarrow> l = 0" | |
| 1475 | by (auto dest!: DERIV_local_max) | |
| 21164 | 1476 | |
| 29975 | 1477 | |
| 60758 | 1478 | subsection \<open>Rolle's Theorem\<close> | 
| 29975 | 1479 | |
| 63558 | 1480 | text \<open>Lemma about introducing open ball in open interval\<close> | 
| 68635 | 1481 | lemma lemma_interval_lt: | 
| 1482 | fixes a b x :: real | |
| 1483 | assumes "a < x" "x < b" | |
| 1484 | shows "\<exists>d. 0 < d \<and> (\<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> a < y \<and> y < b)" | |
| 1485 | using linorder_linear [of "x - a" "b - x"] | |
| 1486 | proof | |
| 1487 | assume "x - a \<le> b - x" | |
| 1488 | with assms show ?thesis | |
| 1489 | by (rule_tac x = "x - a" in exI) auto | |
| 1490 | next | |
| 1491 | assume "b - x \<le> x - a" | |
| 1492 | with assms show ?thesis | |
| 1493 | by (rule_tac x = "b - x" in exI) auto | |
| 1494 | qed | |
| 27668 | 1495 | |
| 63558 | 1496 | lemma lemma_interval: "a < x \<Longrightarrow> x < b \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b)" | 
| 1497 | for a b x :: real | |
| 68635 | 1498 | by (force dest: lemma_interval_lt) | 
| 21164 | 1499 | |
| 63558 | 1500 | text \<open>Rolle's Theorem. | 
| 69593 | 1501 | If \<^term>\<open>f\<close> is defined and continuous on the closed interval | 
| 61799 | 1502 | \<open>[a,b]\<close> and differentiable on the open interval \<open>(a,b)\<close>, | 
| 69593 | 1503 | and \<^term>\<open>f a = f b\<close>, | 
| 1504 | then there exists \<open>x0 \<in> (a,b)\<close> such that \<^term>\<open>f' x0 = 0\<close>\<close> | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1505 | theorem Rolle_deriv: | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1506 | fixes f :: "real \<Rightarrow> real" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1507 | assumes "a < b" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1508 | and fab: "f a = f b" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1509 |     and contf: "continuous_on {a..b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1510 | and derf: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> (f has_derivative f' x) (at x)" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1511 | shows "\<exists>z. a < z \<and> z < b \<and> f' z = (\<lambda>v. 0)" | 
| 21164 | 1512 | proof - | 
| 63558 | 1513 | have le: "a \<le> b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1514 | using \<open>a < b\<close> by simp | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1515 |     have "(a + b) / 2 \<in> {a..b}"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1516 | using assms(1) by auto | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1517 |     then have *: "{a..b} \<noteq> {}"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1518 | by auto | 
| 63558 | 1519 | obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x" and "a \<le> x" "x \<le> b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1520 | using continuous_attains_sup[OF compact_Icc * contf] | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1521 | by (meson atLeastAtMost_iff) | 
| 63558 | 1522 | obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z" and "a \<le> x'" "x' \<le> b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1523 | using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) | 
| 63558 | 1524 | consider "a < x" "x < b" | "x = a \<or> x = b" | 
| 1525 | using \<open>a \<le> x\<close> \<open>x \<le> b\<close> by arith | |
| 1526 | then show ?thesis | |
| 21164 | 1527 | proof cases | 
| 63558 | 1528 | case 1 | 
| 69593 | 1529 | \<comment> \<open>\<^term>\<open>f\<close> attains its maximum within the interval\<close> | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1530 | then obtain l where der: "DERIV f x :> l" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1531 | using derf differentiable_def real_differentiable_def by blast | 
| 63558 | 1532 | obtain d where d: "0 < d" and bound: "\<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | 
| 1533 | using lemma_interval [OF 1] by blast | |
| 1534 | then have bound': "\<forall>y. \<bar>x - y\<bar> < d \<longrightarrow> f y \<le> f x" | |
| 1535 | using x_max by blast | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67399diff
changeset | 1536 | \<comment> \<open>the derivative at a local maximum is zero\<close> | 
| 63558 | 1537 | have "l = 0" | 
| 1538 | by (rule DERIV_local_max [OF der d bound']) | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1539 | with 1 der derf [of x] show ?thesis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1540 | by (metis has_derivative_unique has_field_derivative_def mult_zero_left) | 
| 21164 | 1541 | next | 
| 63558 | 1542 | case 2 | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1543 | then have fx: "f b = f x" by (auto simp add: fab) | 
| 63558 | 1544 | consider "a < x'" "x' < b" | "x' = a \<or> x' = b" | 
| 1545 | using \<open>a \<le> x'\<close> \<open>x' \<le> b\<close> by arith | |
| 1546 | then show ?thesis | |
| 21164 | 1547 | proof cases | 
| 63558 | 1548 | case 1 | 
| 69593 | 1549 | \<comment> \<open>\<^term>\<open>f\<close> attains its minimum within the interval\<close> | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1550 | then obtain l where der: "DERIV f x' :> l" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1551 | using derf differentiable_def real_differentiable_def by blast | 
| 63558 | 1552 | from lemma_interval [OF 1] | 
| 21164 | 1553 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | 
| 63558 | 1554 | by blast | 
| 1555 | then have bound': "\<forall>y. \<bar>x' - y\<bar> < d \<longrightarrow> f x' \<le> f y" | |
| 1556 | using x'_min by blast | |
| 1557 | have "l = 0" by (rule DERIV_local_min [OF der d bound']) | |
| 1558 | \<comment> \<open>the derivative at a local minimum is zero\<close> | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1559 | then show ?thesis using 1 der derf [of x'] | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1560 | by (metis has_derivative_unique has_field_derivative_def mult_zero_left) | 
| 21164 | 1561 | next | 
| 63558 | 1562 | case 2 | 
| 69593 | 1563 | \<comment> \<open>\<^term>\<open>f\<close> is constant throughout the interval\<close> | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1564 | then have fx': "f b = f x'" by (auto simp: fab) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1565 | from dense [OF \<open>a < b\<close>] obtain r where r: "a < r" "r < b" by blast | 
| 63558 | 1566 | obtain d where d: "0 < d" and bound: "\<forall>y. \<bar>r - y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | 
| 1567 | using lemma_interval [OF r] by blast | |
| 1568 | have eq_fb: "f z = f b" if "a \<le> z" and "z \<le> b" for z | |
| 1569 | proof (rule order_antisym) | |
| 1570 | show "f z \<le> f b" by (simp add: fx x_max that) | |
| 1571 | show "f b \<le> f z" by (simp add: fx' x'_min that) | |
| 21164 | 1572 | qed | 
| 63558 | 1573 | have bound': "\<forall>y. \<bar>r - y\<bar> < d \<longrightarrow> f r = f y" | 
| 21164 | 1574 | proof (intro strip) | 
| 63558 | 1575 | fix y :: real | 
| 1576 | assume lt: "\<bar>r - y\<bar> < d" | |
| 1577 | then have "f y = f b" by (simp add: eq_fb bound) | |
| 1578 | then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) | |
| 21164 | 1579 | qed | 
| 63558 | 1580 | obtain l where der: "DERIV f r :> l" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1581 | using derf differentiable_def r(1) r(2) real_differentiable_def by blast | 
| 63558 | 1582 | have "l = 0" | 
| 1583 | by (rule DERIV_local_const [OF der d bound']) | |
| 1584 | \<comment> \<open>the derivative of a constant function is zero\<close> | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1585 | with r der derf [of r] show ?thesis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1586 | by (metis has_derivative_unique has_field_derivative_def mult_zero_left) | 
| 21164 | 1587 | qed | 
| 1588 | qed | |
| 1589 | qed | |
| 1590 | ||
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1591 | corollary Rolle: | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1592 | fixes a b :: real | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1593 |   assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f"
 | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 1594 | and dif [rule_format]: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> f differentiable (at x)" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1595 | shows "\<exists>z. a < z \<and> z < b \<and> DERIV f z :> 0" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1596 | proof - | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1597 | obtain f' where f': "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> (f has_derivative f' x) (at x)" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1598 | using dif unfolding differentiable_def by metis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1599 | then have "\<exists>z. a < z \<and> z < b \<and> f' z = (\<lambda>v. 0)" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1600 | by (metis Rolle_deriv [OF ab]) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1601 | then show ?thesis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1602 | using f' has_derivative_imp_has_field_derivative by fastforce | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1603 | qed | 
| 21164 | 1604 | |
| 63558 | 1605 | subsection \<open>Mean Value Theorem\<close> | 
| 21164 | 1606 | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1607 | theorem mvt: | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1608 | fixes f :: "real \<Rightarrow> real" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1609 | assumes "a < b" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1610 |     and contf: "continuous_on {a..b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1611 | and derf: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> (f has_derivative f' x) (at x)" | 
| 69109 | 1612 | obtains \<xi> where "a < \<xi>" "\<xi> < b" "f b - f a = (f' \<xi>) (b - a)" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1613 | proof - | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1614 | have "\<exists>x. a < x \<and> x < b \<and> (\<lambda>y. f' x y - (f b - f a) / (b - a) * y) = (\<lambda>v. 0)" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1615 | proof (intro Rolle_deriv[OF \<open>a < b\<close>]) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1616 | fix x | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1617 | assume x: "a < x" "x < b" | 
| 69109 | 1618 | show "((\<lambda>x. f x - (f b - f a) / (b - a) * x) | 
| 1619 | has_derivative (\<lambda>y. f' x y - (f b - f a) / (b - a) * y)) (at x)" | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1620 | by (intro derivative_intros derf[OF x]) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1621 | qed (use assms in \<open>auto intro!: continuous_intros simp: field_simps\<close>) | 
| 69109 | 1622 | then obtain \<xi> where | 
| 1623 | "a < \<xi>" "\<xi> < b" "(\<lambda>y. f' \<xi> y - (f b - f a) / (b - a) * y) = (\<lambda>v. 0)" | |
| 1624 | by metis | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1625 | then show ?thesis | 
| 73932 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 desharna parents: 
73795diff
changeset | 1626 | by (metis (no_types, opaque_lifting) that add.right_neutral add_diff_cancel_left' add_diff_eq \<open>a < b\<close> | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1627 | less_irrefl nonzero_eq_divide_eq) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1628 | qed | 
| 21164 | 1629 | |
| 1630 | theorem MVT: | |
| 63558 | 1631 | fixes a b :: real | 
| 1632 | assumes lt: "a < b" | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1633 |     and contf: "continuous_on {a..b} f"
 | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1634 | and dif: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> f differentiable (at x)" | 
| 63558 | 1635 | shows "\<exists>l z. a < z \<and> z < b \<and> DERIV f z :> l \<and> f b - f a = (b - a) * l" | 
| 21164 | 1636 | proof - | 
| 70346 | 1637 | obtain f' :: "real \<Rightarrow> real \<Rightarrow> real" | 
| 1638 | where derf: "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (f has_derivative f' x) (at x)" | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1639 | using dif unfolding differentiable_def by metis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1640 | then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1641 | using mvt [OF lt contf] by blast | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1642 | then show ?thesis | 
| 70346 | 1643 | by (simp add: ac_simps) | 
| 1644 | (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) | |
| 21164 | 1645 | qed | 
| 1646 | ||
| 68635 | 1647 | corollary MVT2: | 
| 1648 | assumes "a < b" and der: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> DERIV f x :> f' x" | |
| 1649 | shows "\<exists>z::real. a < z \<and> z < b \<and> (f b - f a = (b - a) * f' z)" | |
| 1650 | proof - | |
| 1651 | have "\<exists>l z. a < z \<and> | |
| 1652 | z < b \<and> | |
| 1653 | (f has_real_derivative l) (at z) \<and> | |
| 1654 | f b - f a = (b - a) * l" | |
| 1655 | proof (rule MVT [OF \<open>a < b\<close>]) | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1656 |     show "continuous_on {a..b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1657 | by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1658 | show "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> f differentiable (at x)" | 
| 68635 | 1659 | using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) | 
| 1660 | qed | |
| 1661 | with assms show ?thesis | |
| 1662 | by (blast dest: DERIV_unique order_less_imp_le) | |
| 1663 | qed | |
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1664 | |
| 68601 | 1665 | lemma pos_deriv_imp_strict_mono: | 
| 1666 | assumes "\<And>x. (f has_real_derivative f' x) (at x)" | |
| 1667 | assumes "\<And>x. f' x > 0" | |
| 1668 | shows "strict_mono f" | |
| 1669 | proof (rule strict_monoI) | |
| 1670 | fix x y :: real assume xy: "x < y" | |
| 1671 | from assms and xy have "\<exists>z>x. z < y \<and> f y - f x = (y - x) * f' z" | |
| 1672 | by (intro MVT2) (auto dest: connectedD_interval) | |
| 1673 | then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast | |
| 1674 | note \<open>f y - f x = (y - x) * f' z\<close> | |
| 1675 | also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto | |
| 1676 | finally show "f x < f y" by simp | |
| 1677 | qed | |
| 21164 | 1678 | |
| 70614 | 1679 | proposition deriv_nonneg_imp_mono: | 
| 1680 |   assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
 | |
| 1681 |   assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
 | |
| 1682 | assumes ab: "a \<le> b" | |
| 1683 | shows "g a \<le> g b" | |
| 1684 | proof (cases "a < b") | |
| 1685 | assume "a < b" | |
| 1686 | from deriv have "\<And>x. \<lbrakk>x \<ge> a; x \<le> b\<rbrakk> \<Longrightarrow> (g has_real_derivative g' x) (at x)" by simp | |
| 1687 | with MVT2[OF \<open>a < b\<close>] and deriv | |
| 1688 | obtain \<xi> where \<xi>_ab: "\<xi> > a" "\<xi> < b" and g_ab: "g b - g a = (b - a) * g' \<xi>" by blast | |
| 1689 | from \<xi>_ab ab nonneg have "(b - a) * g' \<xi> \<ge> 0" by simp | |
| 1690 | with g_ab show ?thesis by simp | |
| 1691 | qed (insert ab, simp) | |
| 1692 | ||
| 68601 | 1693 | |
| 1694 | subsubsection \<open>A function is constant if its derivative is 0 over an interval.\<close> | |
| 21164 | 1695 | |
| 1696 | lemma DERIV_isconst_end: | |
| 63558 | 1697 | fixes f :: "real \<Rightarrow> real" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1698 |   assumes "a < b" and contf: "continuous_on {a..b} f"
 | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1699 | and 0: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> DERIV f x :> 0" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1700 | shows "f b = f a" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1701 | using MVT [OF \<open>a < b\<close>] "0" DERIV_unique contf real_differentiable_def | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1702 | by (fastforce simp: algebra_simps) | 
| 21164 | 1703 | |
| 1704 | lemma DERIV_isconst2: | |
| 63558 | 1705 | fixes f :: "real \<Rightarrow> real" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1706 |   assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> DERIV f x :> 0"
 | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1707 | and "a \<le> x" "x \<le> b" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1708 | shows "f x = f a" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1709 | proof (cases "a < x") | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1710 | case True | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1711 |   have *: "continuous_on {a..x} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1712 | using \<open>x \<le> b\<close> contf continuous_on_subset by fastforce | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1713 | show ?thesis | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1714 | by (rule DERIV_isconst_end [OF True *]) (use \<open>x \<le> b\<close> derf in auto) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1715 | qed (use \<open>a \<le> x\<close> in auto) | 
| 21164 | 1716 | |
| 63558 | 1717 | lemma DERIV_isconst3: | 
| 1718 | fixes a b x y :: real | |
| 1719 | assumes "a < b" | |
| 1720 |     and "x \<in> {a <..< b}"
 | |
| 1721 |     and "y \<in> {a <..< b}"
 | |
| 1722 |     and derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
 | |
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1723 | 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 | 1724 | 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 | 1725 | case False | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1726 | 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 | 1727 | let ?b = "max x y" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1728 | have *: "DERIV f z :> 0" if "?a \<le> z" "z \<le> ?b" for z | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1729 | proof - | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1730 | have "a < z" and "z < b" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1731 |       using that \<open>x \<in> {a <..< b}\<close> and \<open>y \<in> {a <..< b}\<close> by auto
 | 
| 63558 | 1732 |     then have "z \<in> {a<..<b}" by auto
 | 
| 1733 | then show "DERIV f z :> 0" by (rule derivable) | |
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1734 | qed | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1735 |   have isCont: "continuous_on {?a..?b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1736 | by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1737 | have DERIV: "\<And>z. \<lbrakk>?a < z; z < ?b\<rbrakk> \<Longrightarrow> DERIV f z :> 0" | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1738 | using * by auto | 
| 60758 | 1739 | have "?a < ?b" using \<open>x \<noteq> y\<close> by auto | 
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1740 | 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 | 1741 | 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 | 1742 | qed auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1743 | |
| 21164 | 1744 | lemma DERIV_isconst_all: | 
| 63558 | 1745 | fixes f :: "real \<Rightarrow> real" | 
| 1746 | shows "\<forall>x. DERIV f x :> 0 \<Longrightarrow> f x = f y" | |
| 1747 | apply (rule linorder_cases [of x y]) | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1748 | apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ | 
| 63558 | 1749 | done | 
| 21164 | 1750 | |
| 1751 | lemma DERIV_const_ratio_const: | |
| 63558 | 1752 | fixes f :: "real \<Rightarrow> real" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1753 | assumes "a \<noteq> b" and df: "\<And>x. DERIV f x :> k" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1754 | shows "f b - f a = (b - a) * k" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1755 | proof (cases a b rule: linorder_cases) | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1756 | case less | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1757 | show ?thesis | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1758 | using MVT [OF less] df | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 1759 | by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1760 | next | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1761 | case greater | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1762 | have "f a - f b = (a - b) * k" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1763 | using MVT [OF greater] df | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 1764 | by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1765 | then show ?thesis | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1766 | by (simp add: algebra_simps) | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1767 | qed auto | 
| 21164 | 1768 | |
| 1769 | lemma DERIV_const_ratio_const2: | |
| 63558 | 1770 | fixes f :: "real \<Rightarrow> real" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1771 | assumes "a \<noteq> b" and df: "\<And>x. DERIV f x :> k" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1772 | shows "(f b - f a) / (b - a) = k" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1773 | using DERIV_const_ratio_const [OF assms] \<open>a \<noteq> b\<close> by auto | 
| 21164 | 1774 | |
| 63558 | 1775 | lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" | 
| 1776 | for a b :: real | |
| 1777 | by simp | |
| 21164 | 1778 | |
| 63558 | 1779 | lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" | 
| 1780 | for a b :: real | |
| 1781 | by simp | |
| 21164 | 1782 | |
| 63558 | 1783 | text \<open>Gallileo's "trick": average velocity = av. of end velocities.\<close> | 
| 21164 | 1784 | |
| 1785 | lemma DERIV_const_average: | |
| 63558 | 1786 | fixes v :: "real \<Rightarrow> real" | 
| 1787 | and a b :: real | |
| 1788 | assumes neq: "a \<noteq> b" | |
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1789 | and der: "\<And>x. DERIV v x :> k" | 
| 63558 | 1790 | shows "v ((a + b) / 2) = (v a + v b) / 2" | 
| 21164 | 1791 | proof (cases rule: linorder_cases [of a b]) | 
| 63558 | 1792 | case equal | 
| 1793 | with neq show ?thesis by simp | |
| 21164 | 1794 | next | 
| 1795 | case less | |
| 1796 | have "(v b - v a) / (b - a) = k" | |
| 1797 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 63558 | 1798 | then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" | 
| 1799 | by simp | |
| 21164 | 1800 | moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" | 
| 63558 | 1801 | by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) | 
| 1802 | ultimately show ?thesis | |
| 1803 | using neq by force | |
| 21164 | 1804 | next | 
| 1805 | case greater | |
| 1806 | have "(v b - v a) / (b - a) = k" | |
| 1807 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 63558 | 1808 | then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" | 
| 1809 | by simp | |
| 21164 | 1810 | moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" | 
| 63558 | 1811 | by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) | 
| 1812 | ultimately show ?thesis | |
| 1813 | using neq by (force simp add: add.commute) | |
| 21164 | 1814 | qed | 
| 1815 | ||
| 68601 | 1816 | subsubsection\<open>A function with positive derivative is increasing\<close> | 
| 1817 | text \<open>A simple proof using the MVT, by Jeremy Avigad. And variants.\<close> | |
| 56261 | 1818 | lemma DERIV_pos_imp_increasing_open: | 
| 63558 | 1819 | fixes a b :: real | 
| 1820 | and f :: "real \<Rightarrow> real" | |
| 1821 | assumes "a < b" | |
| 1822 | and "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> (\<exists>y. DERIV f x :> y \<and> y > 0)" | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1823 |     and con: "continuous_on {a..b} f"
 | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1824 | shows "f a < f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1825 | proof (rule ccontr) | 
| 63558 | 1826 | assume f: "\<not> ?thesis" | 
| 1827 | have "\<exists>l z. a < z \<and> z < b \<and> DERIV f z :> l \<and> f b - f a = (b - a) * l" | |
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 1828 | by (rule MVT) (use assms real_differentiable_def in \<open>force+\<close>) | 
| 63558 | 1829 | then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1830 | by auto | 
| 63558 | 1831 | with assms f have "\<not> l > 0" | 
| 36777 
be5461582d0f
avoid using real-specific versions of generic lemmas
 huffman parents: 
35216diff
changeset | 1832 | by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) | 
| 41550 | 1833 | with assms z show False | 
| 56261 | 1834 | by (metis DERIV_unique) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1835 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1836 | |
| 56261 | 1837 | lemma DERIV_pos_imp_increasing: | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1838 | fixes a b :: real and f :: "real \<Rightarrow> real" | 
| 63558 | 1839 | assumes "a < b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1840 | and der: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y > 0" | 
| 56261 | 1841 | shows "f a < f b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1842 | by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \<open>a < b\<close>] der) | 
| 56261 | 1843 | |
| 45791 | 1844 | lemma DERIV_nonneg_imp_nondecreasing: | 
| 63558 | 1845 | fixes a b :: real | 
| 1846 | and f :: "real \<Rightarrow> real" | |
| 1847 | assumes "a \<le> b" | |
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1848 | and "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y \<ge> 0" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1849 | shows "f a \<le> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1850 | proof (rule ccontr, cases "a = b") | 
| 63558 | 1851 | assume "\<not> ?thesis" and "a = b" | 
| 41550 | 1852 | then show False by auto | 
| 37891 | 1853 | next | 
| 63558 | 1854 | assume *: "\<not> ?thesis" | 
| 1855 | assume "a \<noteq> b" | |
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1856 | with \<open>a \<le> b\<close> have "a < b" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1857 | by linarith | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1858 |   moreover have "continuous_on {a..b} f"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1859 | by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1860 | ultimately have "\<exists>l z. a < z \<and> z < b \<and> DERIV f z :> l \<and> f b - f a = (b - a) * l" | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 1861 | using assms MVT [OF \<open>a < b\<close>, of f] real_differentiable_def less_eq_real_def by blast | 
| 63558 | 1862 | then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1863 | by auto | 
| 63558 | 1864 | with * have "a < b" "f b < f a" by auto | 
| 1865 | with ** have "\<not> l \<ge> 0" by (auto simp add: not_le algebra_simps) | |
| 1866 | (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) | |
| 1867 | with assms lz show False | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1868 | by (metis DERIV_unique order_less_imp_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1869 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1870 | |
| 56261 | 1871 | lemma DERIV_neg_imp_decreasing_open: | 
| 63558 | 1872 | fixes a b :: real | 
| 1873 | and f :: "real \<Rightarrow> real" | |
| 1874 | assumes "a < b" | |
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1875 | and "\<And>x. a < x \<Longrightarrow> x < b \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y < 0" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1876 |     and con: "continuous_on {a..b} f"
 | 
| 56261 | 1877 | shows "f a > f b" | 
| 1878 | proof - | |
| 63558 | 1879 | have "(\<lambda>x. -f x) a < (\<lambda>x. -f x) b" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1880 | proof (rule DERIV_pos_imp_increasing_open [of a b]) | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1881 | show "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> \<exists>y. ((\<lambda>x. - f x) has_real_derivative y) (at x) \<and> 0 < y" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1882 | using assms | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1883 | by simp (metis field_differentiable_minus neg_0_less_iff_less) | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1884 |     show "continuous_on {a..b} (\<lambda>x. - f x)"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1885 | using con continuous_on_minus by blast | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1886 | qed (use assms in auto) | 
| 63558 | 1887 | then show ?thesis | 
| 56261 | 1888 | by simp | 
| 1889 | qed | |
| 1890 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1891 | lemma DERIV_neg_imp_decreasing: | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1892 | fixes a b :: real and f :: "real \<Rightarrow> real" | 
| 63558 | 1893 | assumes "a < b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1894 | and der: "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y < 0" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1895 | shows "f a > f b" | 
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1896 | by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \<open>a < b\<close>] der) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1897 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1898 | lemma DERIV_nonpos_imp_nonincreasing: | 
| 63558 | 1899 | fixes a b :: real | 
| 1900 | and f :: "real \<Rightarrow> real" | |
| 1901 | assumes "a \<le> b" | |
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1902 | and "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y \<le> 0" | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1903 | shows "f a \<ge> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1904 | proof - | 
| 63558 | 1905 | have "(\<lambda>x. -f x) a \<le> (\<lambda>x. -f x) b" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1906 | using DERIV_nonneg_imp_nondecreasing [of a b "\<lambda>x. -f x"] assms DERIV_minus by fastforce | 
| 63558 | 1907 | then show ?thesis | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1908 | by simp | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1909 | qed | 
| 21164 | 1910 | |
| 56289 | 1911 | lemma DERIV_pos_imp_increasing_at_bot: | 
| 63558 | 1912 | fixes f :: "real \<Rightarrow> real" | 
| 1913 | assumes "\<And>x. x \<le> b \<Longrightarrow> (\<exists>y. DERIV f x :> y \<and> y > 0)" | |
| 1914 | and lim: "(f \<longlongrightarrow> flim) at_bot" | |
| 56289 | 1915 | shows "flim < f b" | 
| 1916 | proof - | |
| 63952 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1917 | have "\<exists>N. \<forall>n\<le>N. f n \<le> f (b - 1)" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1918 | by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) | 
| 63952 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 1919 | then have "flim \<le> f (b - 1)" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1920 | by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) | 
| 63558 | 1921 | also have "\<dots> < f b" | 
| 56289 | 1922 | by (force intro: DERIV_pos_imp_increasing [where f=f] assms) | 
| 1923 | finally show ?thesis . | |
| 1924 | qed | |
| 1925 | ||
| 1926 | lemma DERIV_neg_imp_decreasing_at_top: | |
| 63558 | 1927 | fixes f :: "real \<Rightarrow> real" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 1928 | assumes der: "\<And>x. x \<ge> b \<Longrightarrow> \<exists>y. DERIV f x :> y \<and> y < 0" | 
| 63558 | 1929 | and lim: "(f \<longlongrightarrow> flim) at_top" | 
| 56289 | 1930 | shows "flim < f b" | 
| 1931 | apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\<lambda>i. f (-i)" and b = "-b", simplified]) | |
| 63558 | 1932 | apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) | 
| 56289 | 1933 | apply (metis filterlim_at_top_mirror lim) | 
| 1934 | done | |
| 1935 | ||
| 60758 | 1936 | text \<open>Derivative of inverse function\<close> | 
| 23041 | 1937 | |
| 1938 | lemma DERIV_inverse_function: | |
| 1939 | fixes f g :: "real \<Rightarrow> real" | |
| 1940 | assumes der: "DERIV f (g x) :> D" | |
| 63558 | 1941 | and neq: "D \<noteq> 0" | 
| 1942 | and x: "a < x" "x < b" | |
| 68611 | 1943 | and inj: "\<And>y. \<lbrakk>a < y; y < b\<rbrakk> \<Longrightarrow> f (g y) = y" | 
| 63558 | 1944 | and cont: "isCont g x" | 
| 23041 | 1945 | shows "DERIV g x :> inverse D" | 
| 68634 | 1946 | unfolding has_field_derivative_iff | 
| 23044 | 1947 | proof (rule LIM_equal2) | 
| 1948 | show "0 < min (x - a) (b - x)" | |
| 63558 | 1949 | using x by arith | 
| 23044 | 1950 | next | 
| 23041 | 1951 | fix y | 
| 23044 | 1952 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 63558 | 1953 | then have "a < y" and "y < b" | 
| 23044 | 1954 | by (simp_all add: abs_less_iff) | 
| 63558 | 1955 | then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" | 
| 23041 | 1956 | by (simp add: inj) | 
| 1957 | next | |
| 61976 | 1958 | have "(\<lambda>z. (f z - f (g x)) / (z - g x)) \<midarrow>g x\<rightarrow> D" | 
| 68634 | 1959 | by (rule der [unfolded has_field_derivative_iff]) | 
| 63558 | 1960 | then have 1: "(\<lambda>z. (f z - x) / (z - g x)) \<midarrow>g x\<rightarrow> D" | 
| 1961 | using inj x by simp | |
| 23041 | 1962 | have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x" | 
| 56219 | 1963 | proof (rule exI, safe) | 
| 23044 | 1964 | show "0 < min (x - a) (b - x)" | 
| 63558 | 1965 | using x by simp | 
| 23041 | 1966 | next | 
| 1967 | fix y | |
| 23044 | 1968 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 63558 | 1969 | then have y: "a < y" "y < b" | 
| 23044 | 1970 | by (simp_all add: abs_less_iff) | 
| 23041 | 1971 | assume "g y = g x" | 
| 63558 | 1972 | then have "f (g y) = f (g x)" by simp | 
| 1973 | then have "y = x" using inj y x by simp | |
| 23041 | 1974 | also assume "y \<noteq> x" | 
| 1975 | finally show False by simp | |
| 1976 | qed | |
| 61976 | 1977 | have "(\<lambda>y. (f (g y) - x) / (g y - g x)) \<midarrow>x\<rightarrow> D" | 
| 23041 | 1978 | using cont 1 2 by (rule isCont_LIM_compose2) | 
| 63558 | 1979 | then show "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x))) \<midarrow>x\<rightarrow> inverse D" | 
| 44568 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44317diff
changeset | 1980 | using neq by (rule tendsto_inverse) | 
| 23041 | 1981 | qed | 
| 1982 | ||
| 60758 | 1983 | subsection \<open>Generalized Mean Value Theorem\<close> | 
| 29975 | 1984 | |
| 21164 | 1985 | theorem GMVT: | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1986 | fixes a b :: real | 
| 21164 | 1987 | assumes alb: "a < b" | 
| 41550 | 1988 | and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1989 | and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable (at x)" | 
| 41550 | 1990 | and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 1991 | and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable (at x)" | 
| 53381 | 1992 | shows "\<exists>g'c f'c c. | 
| 63558 | 1993 | 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 | 1994 | proof - | 
| 63558 | 1995 | let ?h = "\<lambda>x. (f b - f a) * g x - (g b - g a) * f x" | 
| 1996 | have "\<exists>l z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" | |
| 1997 | proof (rule MVT) | |
| 1998 | from assms show "a < b" by simp | |
| 69020 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 1999 |     show "continuous_on {a..b} ?h"
 | 
| 
4f94e262976d
elimination of near duplication involving Rolle's theorem and the MVT
 paulson <lp15@cam.ac.uk> parents: 
68644diff
changeset | 2000 | by (simp add: continuous_at_imp_continuous_on fc gc) | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2001 | show "\<And>x. \<lbrakk>a < x; x < b\<rbrakk> \<Longrightarrow> ?h differentiable (at x)" | 
| 63558 | 2002 | using fd gd by simp | 
| 2003 | qed | |
| 2004 | then obtain l where l: "\<exists>z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 2005 | then obtain c where c: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 21164 | 2006 | |
| 63558 | 2007 | from c have cint: "a < c \<and> c < b" by auto | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 2008 | then obtain g'c where g'c: "DERIV g c :> g'c" | 
| 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 2009 | using gd real_differentiable_def by blast | 
| 63558 | 2010 | from c have "a < c \<and> c < b" by auto | 
| 69022 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 2011 | then obtain f'c where f'c: "DERIV f c :> f'c" | 
| 
e2858770997a
removal of more redundancies, and fixes
 paulson <lp15@cam.ac.uk> parents: 
69020diff
changeset | 2012 | using fd real_differentiable_def by blast | 
| 21164 | 2013 | |
| 63558 | 2014 | from c have "DERIV ?h c :> l" by auto | 
| 41368 | 2015 | moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" | 
| 63558 | 2016 | using g'c f'c by (auto intro!: derivative_eq_intros) | 
| 21164 | 2017 | ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) | 
| 2018 | ||
| 63558 | 2019 | have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" | 
| 2020 | proof - | |
| 2021 | from c have "?h b - ?h a = (b - a) * l" by auto | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 2022 | also from leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp | 
| 63558 | 2023 | finally show ?thesis by simp | 
| 2024 | qed | |
| 2025 | moreover have "?h b - ?h a = 0" | |
| 2026 | proof - | |
| 21164 | 2027 | have "?h b - ?h a = | 
| 63558 | 2028 | ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - | 
| 2029 | ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" | |
| 29667 | 2030 | by (simp add: algebra_simps) | 
| 63558 | 2031 | then show ?thesis by auto | 
| 2032 | qed | |
| 21164 | 2033 | ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto | 
| 2034 | with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp | |
| 63558 | 2035 | then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp | 
| 2036 | then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) | |
| 2037 | with g'c f'c cint show ?thesis by auto | |
| 21164 | 2038 | qed | 
| 2039 | ||
| 50327 | 2040 | lemma GMVT': | 
| 2041 | fixes f g :: "real \<Rightarrow> real" | |
| 2042 | assumes "a < b" | |
| 63558 | 2043 | and isCont_f: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont f z" | 
| 2044 | and isCont_g: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont g z" | |
| 2045 | and DERIV_g: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV g z :> (g' z)" | |
| 2046 | and DERIV_f: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV f z :> (f' z)" | |
| 50327 | 2047 | shows "\<exists>c. a < c \<and> c < b \<and> (f b - f a) * g' c = (g b - g a) * f' c" | 
| 2048 | proof - | |
| 2049 | have "\<exists>g'c f'c c. DERIV g c :> g'c \<and> DERIV f c :> f'c \<and> | |
| 63558 | 2050 | a < c \<and> c < b \<and> (f b - f a) * g'c = (g b - g a) * f'c" | 
| 56181 
2aa0b19e74f3
unify syntax for has_derivative and differentiable
 hoelzl parents: 
55970diff
changeset | 2051 | using assms by (intro GMVT) (force simp: real_differentiable_def)+ | 
| 50327 | 2052 | then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" | 
| 2053 | using DERIV_f DERIV_g by (force dest: DERIV_unique) | |
| 2054 | then show ?thesis | |
| 2055 | by auto | |
| 2056 | qed | |
| 2057 | ||
| 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 | 2058 | |
| 60758 | 2059 | subsection \<open>L'Hopitals rule\<close> | 
| 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 | 2060 | |
| 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 | 2061 | 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 | 2062 | fixes a :: "'a :: linorder_topology" | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2063 | assumes "continuous (at_left a) g" and f: "(f \<longlongrightarrow> g a) (at_right a)" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2064 | shows "isCont (\<lambda>x. if x \<le> a then g x else f x) a" (is "isCont ?gf a") | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2065 | proof - | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2066 | have g: "(g \<longlongrightarrow> g a) (at_left a)" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2067 | using assms continuous_within by blast | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2068 | show ?thesis | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2069 | unfolding isCont_def continuous_within | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2070 | proof (intro filterlim_split_at; simp) | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2071 | show "(?gf \<longlongrightarrow> g a) (at_left a)" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2072 | by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2073 | show "(?gf \<longlongrightarrow> g a) (at_right a)" | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2074 | by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2075 | qed | 
| 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2076 | qed | 
| 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 | 2077 | |
| 50327 | 2078 | lemma lhopital_right_0: | 
| 50329 | 2079 | fixes f0 g0 :: "real \<Rightarrow> real" | 
| 61973 | 2080 | assumes f_0: "(f0 \<longlongrightarrow> 0) (at_right 0)" | 
| 63558 | 2081 | and g_0: "(g0 \<longlongrightarrow> 0) (at_right 0)" | 
| 2082 | and ev: | |
| 2083 | "eventually (\<lambda>x. g0 x \<noteq> 0) (at_right 0)" | |
| 2084 | "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)" | |
| 2085 | "eventually (\<lambda>x. DERIV f0 x :> f' x) (at_right 0)" | |
| 2086 | "eventually (\<lambda>x. DERIV g0 x :> g' x) (at_right 0)" | |
| 63713 | 2087 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) F (at_right 0)" | 
| 2088 | shows "filterlim (\<lambda> x. f0 x / g0 x) F (at_right 0)" | |
| 50327 | 2089 | proof - | 
| 63040 | 2090 | define f where [abs_def]: "f x = (if x \<le> 0 then 0 else f0 x)" for x | 
| 50329 | 2091 | then have "f 0 = 0" by simp | 
| 2092 | ||
| 63040 | 2093 | define g where [abs_def]: "g x = (if x \<le> 0 then 0 else g0 x)" for x | 
| 50329 | 2094 | then have "g 0 = 0" by simp | 
| 2095 | ||
| 2096 | have "eventually (\<lambda>x. g0 x \<noteq> 0 \<and> g' x \<noteq> 0 \<and> | |
| 2097 | DERIV f0 x :> (f' x) \<and> DERIV g0 x :> (g' x)) (at_right 0)" | |
| 2098 | using ev by eventually_elim auto | |
| 2099 | then obtain a where [arith]: "0 < a" | |
| 2100 | and g0_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g0 x \<noteq> 0" | |
| 50327 | 2101 | and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0" | 
| 50329 | 2102 | and f0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV f0 x :> (f' x)" | 
| 2103 | and g0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV g0 x :> (g' x)" | |
| 56219 | 2104 | unfolding eventually_at by (auto simp: dist_real_def) | 
| 50327 | 2105 | |
| 50329 | 2106 | have g_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g x \<noteq> 0" | 
| 2107 | using g0_neq_0 by (simp add: g_def) | |
| 2108 | ||
| 63558 | 2109 | have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x | 
| 2110 | using that | |
| 2111 | by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) | |
| 2112 | (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) | |
| 50329 | 2113 | |
| 63558 | 2114 | have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x | 
| 2115 | using that | |
| 2116 | by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) | |
| 2117 | (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) | |
| 50329 | 2118 | |
| 2119 | 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 | 2120 | 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 | 2121 | |
| 50329 | 2122 | 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 | 2123 | unfolding g_def by (intro isCont_If_ge g_0 continuous_const) | 
| 50329 | 2124 | |
| 50327 | 2125 |   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)"
 | 
| 63558 | 2126 | proof (rule bchoice, rule ballI) | 
| 2127 | fix x | |
| 2128 |     assume "x \<in> {0 <..< a}"
 | |
| 50327 | 2129 | then have x[arith]: "0 < x" "x < a" by auto | 
| 60758 | 2130 | with g'_neq_0 g_neq_0 \<open>g 0 = 0\<close> have g': "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x" | 
| 50327 | 2131 | by auto | 
| 50328 | 2132 | have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont f x" | 
| 60758 | 2133 | using \<open>isCont f 0\<close> f by (auto intro: DERIV_isCont simp: le_less) | 
| 50328 | 2134 | moreover have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont g x" | 
| 60758 | 2135 | using \<open>isCont g 0\<close> g by (auto intro: DERIV_isCont simp: le_less) | 
| 50328 | 2136 | ultimately have "\<exists>c. 0 < c \<and> c < x \<and> (f x - f 0) * g' c = (g x - g 0) * f' c" | 
| 60758 | 2137 | using f g \<open>x < a\<close> by (intro GMVT') auto | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 2138 | 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 | 2139 | by blast | 
| 50327 | 2140 | moreover | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
51642diff
changeset | 2141 | from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" | 
| 50327 | 2142 | by (simp add: field_simps) | 
| 2143 | ultimately show "\<exists>y. 0 < y \<and> y < x \<and> f x / g x = f' y / g' y" | |
| 60758 | 2144 | using \<open>f 0 = 0\<close> \<open>g 0 = 0\<close> by (auto intro!: exI[of _ c]) | 
| 50327 | 2145 | qed | 
| 53381 | 2146 |   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 | 2147 | 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 | 2148 | unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) | 
| 50327 | 2149 | moreover | 
| 2150 | from \<zeta> have "eventually (\<lambda>x. norm (\<zeta> x) \<le> x) (at_right 0)" | |
| 2151 | by eventually_elim auto | |
| 61973 | 2152 | then have "((\<lambda>x. norm (\<zeta> x)) \<longlongrightarrow> 0) (at_right 0)" | 
| 58729 
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
 hoelzl parents: 
57953diff
changeset | 2153 | by (rule_tac real_tendsto_sandwich[where f="\<lambda>x. 0" and h="\<lambda>x. x"]) auto | 
| 61973 | 2154 | then have "(\<zeta> \<longlongrightarrow> 0) (at_right 0)" | 
| 50327 | 2155 | by (rule tendsto_norm_zero_cancel) | 
| 2156 | with \<zeta> have "filterlim \<zeta> (at_right 0) (at_right 0)" | |
| 61810 | 2157 | by (auto elim!: eventually_mono simp: filterlim_at) | 
| 63713 | 2158 | from this lim have "filterlim (\<lambda>t. f' (\<zeta> t) / g' (\<zeta> t)) F (at_right 0)" | 
| 50327 | 2159 | by (rule_tac filterlim_compose[of _ _ _ \<zeta>]) | 
| 63713 | 2160 | ultimately have "filterlim (\<lambda>t. f t / g t) F (at_right 0)" (is ?P) | 
| 50328 | 2161 | by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) | 
| 61810 | 2162 | (auto elim: eventually_mono) | 
| 50329 | 2163 | 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 | 2164 | by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) | 
| 50329 | 2165 | finally show ?thesis . | 
| 50327 | 2166 | qed | 
| 2167 | ||
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2168 | lemma lhopital_right: | 
| 63558 | 2169 | "(f \<longlongrightarrow> 0) (at_right x) \<Longrightarrow> (g \<longlongrightarrow> 0) (at_right x) \<Longrightarrow> | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2170 | 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 | 2171 | 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 | 2172 | 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 | 2173 | eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow> | 
| 63713 | 2174 | filterlim (\<lambda> x. (f' x / g' x)) F (at_right x) \<Longrightarrow> | 
| 2175 | filterlim (\<lambda> x. f x / g x) F (at_right x)" | |
| 63558 | 2176 | for x :: real | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2177 | 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 | 2178 | 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 | 2179 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2180 | lemma lhopital_left: | 
| 63558 | 2181 | "(f \<longlongrightarrow> 0) (at_left x) \<Longrightarrow> (g \<longlongrightarrow> 0) (at_left x) \<Longrightarrow> | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2182 | 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 | 2183 | 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 | 2184 | 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 | 2185 | eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow> | 
| 63713 | 2186 | filterlim (\<lambda> x. (f' x / g' x)) F (at_left x) \<Longrightarrow> | 
| 2187 | filterlim (\<lambda> x. f x / g x) F (at_left x)" | |
| 63558 | 2188 | for x :: real | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2189 | unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56409diff
changeset | 2190 | by (rule lhopital_right[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror) | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2191 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2192 | lemma lhopital: | 
| 63558 | 2193 | "(f \<longlongrightarrow> 0) (at x) \<Longrightarrow> (g \<longlongrightarrow> 0) (at x) \<Longrightarrow> | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2194 | 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 | 2195 | 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 | 2196 | 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 | 2197 | eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow> | 
| 63713 | 2198 | filterlim (\<lambda> x. (f' x / g' x)) F (at x) \<Longrightarrow> | 
| 2199 | filterlim (\<lambda> x. f x / g x) F (at x)" | |
| 63558 | 2200 | for x :: real | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2201 | 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 | 2202 | 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 | 2203 | |
| 63713 | 2204 | |
| 50327 | 2205 | lemma lhopital_right_0_at_top: | 
| 2206 | fixes f g :: "real \<Rightarrow> real" | |
| 2207 | assumes g_0: "LIM x at_right 0. g x :> at_top" | |
| 63558 | 2208 | and ev: | 
| 2209 | "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)" | |
| 2210 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_right 0)" | |
| 2211 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_right 0)" | |
| 2212 | and lim: "((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> x) (at_right 0)" | |
| 61973 | 2213 | shows "((\<lambda> x. f x / g x) \<longlongrightarrow> x) (at_right 0)" | 
| 50327 | 2214 | unfolding tendsto_iff | 
| 2215 | proof safe | |
| 63558 | 2216 | fix e :: real | 
| 2217 | assume "0 < e" | |
| 50327 | 2218 | with lim[unfolded tendsto_iff, rule_format, of "e / 4"] | 
| 63558 | 2219 | have "eventually (\<lambda>t. dist (f' t / g' t) x < e / 4) (at_right 0)" | 
| 2220 | by simp | |
| 50327 | 2221 | from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] | 
| 2222 | obtain a where [arith]: "0 < a" | |
| 2223 | and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0" | |
| 2224 | and f0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV f x :> (f' x)" | |
| 2225 | and g0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV g x :> (g' x)" | |
| 2226 | 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 | 2227 | unfolding eventually_at_le by (auto simp: dist_real_def) | 
| 50327 | 2228 | |
| 63558 | 2229 | from Df have "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 | 2230 | unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) | 
| 50327 | 2231 | |
| 2232 | moreover | |
| 50328 | 2233 | have "eventually (\<lambda>t. 0 < g t) (at_right 0)" "eventually (\<lambda>t. g a < g t) (at_right 0)" | 
| 61810 | 2234 | using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) | 
| 50327 | 2235 | |
| 2236 | moreover | |
| 61973 | 2237 | have inv_g: "((\<lambda>x. inverse (g x)) \<longlongrightarrow> 0) (at_right 0)" | 
| 50327 | 2238 | using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] | 
| 2239 | by (rule filterlim_compose) | |
| 61973 | 2240 | then have "((\<lambda>x. norm (1 - g a * inverse (g x))) \<longlongrightarrow> norm (1 - g a * 0)) (at_right 0)" | 
| 50327 | 2241 | by (intro tendsto_intros) | 
| 61973 | 2242 | then have "((\<lambda>x. norm (1 - g a / g x)) \<longlongrightarrow> 1) (at_right 0)" | 
| 50327 | 2243 | by (simp add: inverse_eq_divide) | 
| 2244 | from this[unfolded tendsto_iff, rule_format, of 1] | |
| 2245 | have "eventually (\<lambda>x. norm (1 - g a / g x) < 2) (at_right 0)" | |
| 61810 | 2246 | by (auto elim!: eventually_mono simp: dist_real_def) | 
| 50327 | 2247 | |
| 2248 | moreover | |
| 63558 | 2249 | from inv_g have "((\<lambda>t. norm ((f a - x * g a) * inverse (g t))) \<longlongrightarrow> norm ((f a - x * g a) * 0)) | 
| 2250 | (at_right 0)" | |
| 50327 | 2251 | by (intro tendsto_intros) | 
| 61973 | 2252 | then have "((\<lambda>t. norm (f a - x * g a) / norm (g t)) \<longlongrightarrow> 0) (at_right 0)" | 
| 50327 | 2253 | by (simp add: inverse_eq_divide) | 
| 60758 | 2254 | from this[unfolded tendsto_iff, rule_format, of "e / 2"] \<open>0 < e\<close> | 
| 50327 | 2255 | have "eventually (\<lambda>t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" | 
| 2256 | by (auto simp: dist_real_def) | |
| 2257 | ||
| 2258 | ultimately show "eventually (\<lambda>t. dist (f t / g t) x < e) (at_right 0)" | |
| 2259 | proof eventually_elim | |
| 2260 | fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" | |
| 2261 | assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" | |
| 2262 | ||
| 2263 | have "\<exists>y. t < y \<and> y < a \<and> (g a - g t) * f' y = (f a - f t) * g' y" | |
| 2264 | using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ | |
| 53381 | 2265 | then obtain y where [arith]: "t < y" "y < a" | 
| 2266 | and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" | |
| 2267 | by blast | |
| 2268 | from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" | |
| 60758 | 2269 | using \<open>g a < g t\<close> g'_neq_0[of y] by (auto simp add: field_simps) | 
| 50327 | 2270 | |
| 2271 | 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" | |
| 2272 | by (simp add: field_simps) | |
| 2273 | have "norm (f t / g t - x) \<le> | |
| 2274 | norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" | |
| 2275 | unfolding * by (rule norm_triangle_ineq) | |
| 2276 | also have "\<dots> = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" | |
| 2277 | by (simp add: abs_mult D_eq dist_real_def) | |
| 2278 | also have "\<dots> < (e / 4) * 2 + e / 2" | |
| 60758 | 2279 | using ineq Df[of y] \<open>0 < e\<close> by (intro add_le_less_mono mult_mono) auto | 
| 50327 | 2280 | finally show "dist (f t / g t) x < e" | 
| 2281 | by (simp add: dist_real_def) | |
| 2282 | qed | |
| 2283 | qed | |
| 2284 | ||
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2285 | 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 | 2286 | "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 | 2287 | 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 | 2288 | 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 | 2289 | eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow> | 
| 61973 | 2290 | ((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> y) (at_right x) \<Longrightarrow> | 
| 2291 | ((\<lambda> x. f x / g x) \<longlongrightarrow> y) (at_right x)" | |
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2292 | 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 | 2293 | 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 | 2294 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2295 | lemma lhopital_left_at_top: | 
| 63558 | 2296 | "LIM x at_left x. g x :> at_top \<Longrightarrow> | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2297 | 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 | 2298 | 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 | 2299 | eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow> | 
| 61973 | 2300 | ((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> y) (at_left x) \<Longrightarrow> | 
| 2301 | ((\<lambda> x. f x / g x) \<longlongrightarrow> y) (at_left x)" | |
| 63558 | 2302 | for x :: real | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2303 | unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56409diff
changeset | 2304 | by (rule lhopital_right_at_top[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror) | 
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2305 | |
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2306 | lemma lhopital_at_top: | 
| 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2307 | "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 | 2308 | 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 | 2309 | 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 | 2310 | eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow> | 
| 61973 | 2311 | ((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> y) (at x) \<Longrightarrow> | 
| 2312 | ((\<lambda> x. f x / g x) \<longlongrightarrow> y) (at x)" | |
| 50330 
d0b12171118e
conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
 hoelzl parents: 
50329diff
changeset | 2313 | 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 | 2314 | 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 | 2315 | |
| 50347 | 2316 | lemma lhospital_at_top_at_top: | 
| 2317 | fixes f g :: "real \<Rightarrow> real" | |
| 2318 | assumes g_0: "LIM x at_top. g x :> at_top" | |
| 63558 | 2319 | and g': "eventually (\<lambda>x. g' x \<noteq> 0) at_top" | 
| 2320 | and Df: "eventually (\<lambda>x. DERIV f x :> f' x) at_top" | |
| 2321 | and Dg: "eventually (\<lambda>x. DERIV g x :> g' x) at_top" | |
| 2322 | and lim: "((\<lambda> x. (f' x / g' x)) \<longlongrightarrow> x) at_top" | |
| 61973 | 2323 | shows "((\<lambda> x. f x / g x) \<longlongrightarrow> x) at_top" | 
| 50347 | 2324 | unfolding filterlim_at_top_to_right | 
| 2325 | proof (rule lhopital_right_0_at_top) | |
| 2326 | let ?F = "\<lambda>x. f (inverse x)" | |
| 2327 | let ?G = "\<lambda>x. g (inverse x)" | |
| 2328 | let ?R = "at_right (0::real)" | |
| 2329 | let ?D = "\<lambda>f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" | |
| 2330 | show "LIM x ?R. ?G x :> at_top" | |
| 2331 | using g_0 unfolding filterlim_at_top_to_right . | |
| 2332 | show "eventually (\<lambda>x. DERIV ?G x :> ?D g' x) ?R" | |
| 2333 | unfolding eventually_at_right_to_top | |
| 63558 | 2334 | using Dg eventually_ge_at_top[where c=1] | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2335 | by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ | 
| 50347 | 2336 | show "eventually (\<lambda>x. DERIV ?F x :> ?D f' x) ?R" | 
| 2337 | unfolding eventually_at_right_to_top | |
| 63558 | 2338 | using Df eventually_ge_at_top[where c=1] | 
| 68638 
87d1bff264df
de-applying and meta-quantifying
 paulson <lp15@cam.ac.uk> parents: 
68635diff
changeset | 2339 | by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ | 
| 50347 | 2340 | show "eventually (\<lambda>x. ?D g' x \<noteq> 0) ?R" | 
| 2341 | unfolding eventually_at_right_to_top | |
| 63558 | 2342 | using g' eventually_ge_at_top[where c=1] | 
| 50347 | 2343 | by eventually_elim auto | 
| 61973 | 2344 | show "((\<lambda>x. ?D f' x / ?D g' x) \<longlongrightarrow> x) ?R" | 
| 50347 | 2345 | unfolding filterlim_at_right_to_top | 
| 2346 | apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) | |
| 63558 | 2347 | using eventually_ge_at_top[where c=1] | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56409diff
changeset | 2348 | by eventually_elim simp | 
| 50347 | 2349 | qed | 
| 2350 | ||
| 63713 | 2351 | lemma lhopital_right_at_top_at_top: | 
| 2352 | fixes f g :: "real \<Rightarrow> real" | |
| 2353 | assumes f_0: "LIM x at_right a. f x :> at_top" | |
| 2354 | assumes g_0: "LIM x at_right a. g x :> at_top" | |
| 2355 | and ev: | |
| 2356 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_right a)" | |
| 2357 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_right a)" | |
| 2358 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_top (at_right a)" | |
| 2359 | shows "filterlim (\<lambda> x. f x / g x) at_top (at_right a)" | |
| 2360 | proof - | |
| 2361 | from lim have pos: "eventually (\<lambda>x. f' x / g' x > 0) (at_right a)" | |
| 2362 | unfolding filterlim_at_top_dense by blast | |
| 2363 | have "((\<lambda>x. g x / f x) \<longlongrightarrow> 0) (at_right a)" | |
| 2364 | proof (rule lhopital_right_at_top) | |
| 2365 | from pos show "eventually (\<lambda>x. f' x \<noteq> 0) (at_right a)" by eventually_elim auto | |
| 2366 | from tendsto_inverse_0_at_top[OF lim] | |
| 2367 | show "((\<lambda>x. g' x / f' x) \<longlongrightarrow> 0) (at_right a)" by simp | |
| 2368 | qed fact+ | |
| 2369 | moreover from f_0 g_0 | |
| 2370 | have "eventually (\<lambda>x. f x > 0) (at_right a)" "eventually (\<lambda>x. g x > 0) (at_right a)" | |
| 2371 | unfolding filterlim_at_top_dense by blast+ | |
| 2372 | hence "eventually (\<lambda>x. g x / f x > 0) (at_right a)" by eventually_elim simp | |
| 2373 | ultimately have "filterlim (\<lambda>x. inverse (g x / f x)) at_top (at_right a)" | |
| 2374 | by (rule filterlim_inverse_at_top) | |
| 2375 | thus ?thesis by simp | |
| 2376 | qed | |
| 63717 | 2377 | |
| 63713 | 2378 | lemma lhopital_right_at_top_at_bot: | 
| 2379 | fixes f g :: "real \<Rightarrow> real" | |
| 2380 | assumes f_0: "LIM x at_right a. f x :> at_top" | |
| 2381 | assumes g_0: "LIM x at_right a. g x :> at_bot" | |
| 2382 | and ev: | |
| 2383 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_right a)" | |
| 2384 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_right a)" | |
| 2385 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_bot (at_right a)" | |
| 2386 | shows "filterlim (\<lambda> x. f x / g x) at_bot (at_right a)" | |
| 2387 | proof - | |
| 2388 | from ev(2) have ev': "eventually (\<lambda>x. DERIV (\<lambda>x. -g x) x :> -g' x) (at_right a)" | |
| 2389 | by eventually_elim (auto intro: derivative_intros) | |
| 2390 | have "filterlim (\<lambda>x. f x / (-g x)) at_top (at_right a)" | |
| 2391 | by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\<lambda>x. -g' x"]) | |
| 2392 | (insert assms ev', auto simp: filterlim_uminus_at_bot) | |
| 2393 | hence "filterlim (\<lambda>x. -(f x / g x)) at_top (at_right a)" by simp | |
| 2394 | thus ?thesis by (simp add: filterlim_uminus_at_bot) | |
| 2395 | qed | |
| 2396 | ||
| 2397 | lemma lhopital_left_at_top_at_top: | |
| 2398 | fixes f g :: "real \<Rightarrow> real" | |
| 2399 | assumes f_0: "LIM x at_left a. f x :> at_top" | |
| 2400 | assumes g_0: "LIM x at_left a. g x :> at_top" | |
| 2401 | and ev: | |
| 2402 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_left a)" | |
| 2403 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_left a)" | |
| 2404 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_top (at_left a)" | |
| 2405 | shows "filterlim (\<lambda> x. f x / g x) at_top (at_left a)" | |
| 2406 | by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, | |
| 2407 | rule lhopital_right_at_top_at_top[where f'="\<lambda>x. - f' (- x)"]) | |
| 2408 | (insert assms, auto simp: DERIV_mirror) | |
| 2409 | ||
| 2410 | lemma lhopital_left_at_top_at_bot: | |
| 2411 | fixes f g :: "real \<Rightarrow> real" | |
| 2412 | assumes f_0: "LIM x at_left a. f x :> at_top" | |
| 2413 | assumes g_0: "LIM x at_left a. g x :> at_bot" | |
| 2414 | and ev: | |
| 2415 | "eventually (\<lambda>x. DERIV f x :> f' x) (at_left a)" | |
| 2416 | "eventually (\<lambda>x. DERIV g x :> g' x) (at_left a)" | |
| 2417 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_bot (at_left a)" | |
| 2418 | shows "filterlim (\<lambda> x. f x / g x) at_bot (at_left a)" | |
| 2419 | by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, | |
| 2420 | rule lhopital_right_at_top_at_bot[where f'="\<lambda>x. - f' (- x)"]) | |
| 2421 | (insert assms, auto simp: DERIV_mirror) | |
| 2422 | ||
| 2423 | lemma lhopital_at_top_at_top: | |
| 2424 | fixes f g :: "real \<Rightarrow> real" | |
| 2425 | assumes f_0: "LIM x at a. f x :> at_top" | |
| 2426 | assumes g_0: "LIM x at a. g x :> at_top" | |
| 2427 | and ev: | |
| 2428 | "eventually (\<lambda>x. DERIV f x :> f' x) (at a)" | |
| 2429 | "eventually (\<lambda>x. DERIV g x :> g' x) (at a)" | |
| 2430 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_top (at a)" | |
| 2431 | shows "filterlim (\<lambda> x. f x / g x) at_top (at a)" | |
| 2432 | using assms unfolding eventually_at_split filterlim_at_split | |
| 2433 | by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] | |
| 2434 | lhopital_left_at_top_at_top[of f a g f' g']) | |
| 2435 | ||
| 2436 | lemma lhopital_at_top_at_bot: | |
| 2437 | fixes f g :: "real \<Rightarrow> real" | |
| 2438 | assumes f_0: "LIM x at a. f x :> at_top" | |
| 2439 | assumes g_0: "LIM x at a. g x :> at_bot" | |
| 2440 | and ev: | |
| 2441 | "eventually (\<lambda>x. DERIV f x :> f' x) (at a)" | |
| 2442 | "eventually (\<lambda>x. DERIV g x :> g' x) (at a)" | |
| 2443 | and lim: "filterlim (\<lambda> x. (f' x / g' x)) at_bot (at a)" | |
| 2444 | shows "filterlim (\<lambda> x. f x / g x) at_bot (at a)" | |
| 2445 | using assms unfolding eventually_at_split filterlim_at_split | |
| 2446 | by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] | |
| 2447 | lhopital_left_at_top_at_bot[of f a g f' g']) | |
| 2448 | ||
| 21164 | 2449 | end |