| author | nipkow | 
| Thu, 01 Apr 2010 09:31:58 +0200 | |
| changeset 36070 | d80e5d3c8fe1 | 
| parent 35216 | 7641e8d831d2 | 
| child 36777 | be5461582d0f | 
| permissions | -rw-r--r-- | 
| 21164 | 1 | (* Title : Deriv.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | |
| 5 | GMVT by Benjamin Porter, 2005 | |
| 6 | *) | |
| 7 | ||
| 8 | header{* Differentiation *}
 | |
| 9 | ||
| 10 | theory Deriv | |
| 29987 | 11 | imports Lim | 
| 21164 | 12 | begin | 
| 13 | ||
| 22984 | 14 | text{*Standard Definitions*}
 | 
| 21164 | 15 | |
| 16 | definition | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 17 | deriv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool" | 
| 21164 | 18 |     --{*Differentiation: D is derivative of function f at x*}
 | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21239diff
changeset | 19 |           ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 20 | "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)" | 
| 21164 | 21 | |
| 22 | primrec | |
| 34941 | 23 | Bolzano_bisect :: "(real \<times> real \<Rightarrow> bool) \<Rightarrow> real \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real \<times> real" where | 
| 24 | "Bolzano_bisect P a b 0 = (a, b)" | |
| 25 | | "Bolzano_bisect P a b (Suc n) = | |
| 26 | (let (x, y) = Bolzano_bisect P a b n | |
| 27 | in if P (x, (x+y) / 2) then ((x+y)/2, y) | |
| 28 | else (x, (x+y)/2))" | |
| 21164 | 29 | |
| 30 | ||
| 31 | subsection {* Derivatives *}
 | |
| 32 | ||
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 33 | lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)" | 
| 21164 | 34 | by (simp add: deriv_def) | 
| 35 | ||
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 36 | lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D" | 
| 21164 | 37 | by (simp add: deriv_def) | 
| 38 | ||
| 39 | lemma DERIV_const [simp]: "DERIV (\<lambda>x. k) x :> 0" | |
| 40 | by (simp add: deriv_def) | |
| 41 | ||
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23044diff
changeset | 42 | lemma DERIV_ident [simp]: "DERIV (\<lambda>x. x) x :> 1" | 
| 23398 | 43 | by (simp add: deriv_def cong: LIM_cong) | 
| 21164 | 44 | |
| 45 | lemma add_diff_add: | |
| 46 | fixes a b c d :: "'a::ab_group_add" | |
| 47 | shows "(a + c) - (b + d) = (a - b) + (c - d)" | |
| 48 | by simp | |
| 49 | ||
| 50 | lemma DERIV_add: | |
| 51 | "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + g x) x :> D + E" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 52 | by (simp only: deriv_def add_diff_add add_divide_distrib LIM_add) | 
| 21164 | 53 | |
| 54 | lemma DERIV_minus: | |
| 55 | "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. - f x) x :> - D" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 56 | by (simp only: deriv_def minus_diff_minus divide_minus_left LIM_minus) | 
| 21164 | 57 | |
| 58 | lemma DERIV_diff: | |
| 59 | "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x - g x) x :> D - E" | |
| 60 | by (simp only: diff_def DERIV_add DERIV_minus) | |
| 61 | ||
| 62 | lemma DERIV_add_minus: | |
| 63 | "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + - g x) x :> D + - E" | |
| 64 | by (simp only: DERIV_add DERIV_minus) | |
| 65 | ||
| 66 | lemma DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x" | |
| 67 | proof (unfold isCont_iff) | |
| 68 | assume "DERIV f x :> D" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 69 | hence "(\<lambda>h. (f(x+h) - f(x)) / h) -- 0 --> D" | 
| 21164 | 70 | by (rule DERIV_D) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 71 | hence "(\<lambda>h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0" | 
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23044diff
changeset | 72 | by (intro LIM_mult LIM_ident) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 73 | hence "(\<lambda>h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 74 | by simp | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 75 | hence "(\<lambda>h. f(x+h) - f(x)) -- 0 --> 0" | 
| 23398 | 76 | by (simp cong: LIM_cong) | 
| 21164 | 77 | thus "(\<lambda>h. f(x+h)) -- 0 --> f(x)" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 78 | by (simp add: LIM_def dist_norm) | 
| 21164 | 79 | qed | 
| 80 | ||
| 81 | lemma DERIV_mult_lemma: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 82 | fixes a b c d :: "'a::real_field" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 83 | shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23441diff
changeset | 84 | by (simp add: diff_minus add_divide_distrib [symmetric] ring_distribs) | 
| 21164 | 85 | |
| 86 | lemma DERIV_mult': | |
| 87 | assumes f: "DERIV f x :> D" | |
| 88 | assumes g: "DERIV g x :> E" | |
| 89 | shows "DERIV (\<lambda>x. f x * g x) x :> f x * E + D * g x" | |
| 90 | proof (unfold deriv_def) | |
| 91 | from f have "isCont f x" | |
| 92 | by (rule DERIV_isCont) | |
| 93 | hence "(\<lambda>h. f(x+h)) -- 0 --> f x" | |
| 94 | by (simp only: isCont_iff) | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 95 | hence "(\<lambda>h. f(x+h) * ((g(x+h) - g x) / h) + | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 96 | ((f(x+h) - f x) / h) * g x) | 
| 21164 | 97 | -- 0 --> f x * E + D * g x" | 
| 22613 | 98 | by (intro LIM_add LIM_mult LIM_const DERIV_D f g) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 99 | thus "(\<lambda>h. (f(x+h) * g(x+h) - f x * g x) / h) | 
| 21164 | 100 | -- 0 --> f x * E + D * g x" | 
| 101 | by (simp only: DERIV_mult_lemma) | |
| 102 | qed | |
| 103 | ||
| 104 | lemma DERIV_mult: | |
| 105 | "[| DERIV f x :> Da; DERIV g x :> Db |] | |
| 106 | ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" | |
| 107 | by (drule (1) DERIV_mult', simp only: mult_commute add_commute) | |
| 108 | ||
| 109 | lemma DERIV_unique: | |
| 110 | "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E" | |
| 111 | apply (simp add: deriv_def) | |
| 112 | apply (blast intro: LIM_unique) | |
| 113 | done | |
| 114 | ||
| 115 | text{*Differentiation of finite sum*}
 | |
| 116 | ||
| 31880 | 117 | lemma DERIV_setsum: | 
| 118 | assumes "finite S" | |
| 119 | and "\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x :> (f' x n)" | |
| 120 | shows "DERIV (%x. setsum (f x) S) x :> setsum (f' x) S" | |
| 121 | using assms by induct (auto intro!: DERIV_add) | |
| 122 | ||
| 21164 | 123 | lemma DERIV_sumr [rule_format (no_asm)]: | 
| 124 | "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x)) | |
| 125 | --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x :> (\<Sum>r=m..<n. f' r x)" | |
| 31880 | 126 | by (auto intro: DERIV_setsum) | 
| 21164 | 127 | |
| 128 | text{*Alternative definition for differentiability*}
 | |
| 129 | ||
| 130 | lemma DERIV_LIM_iff: | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 131 |   fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 132 | "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) = | 
| 21164 | 133 | ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)" | 
| 134 | apply (rule iffI) | |
| 135 | apply (drule_tac k="- a" in LIM_offset) | |
| 136 | apply (simp add: diff_minus) | |
| 137 | apply (drule_tac k="a" in LIM_offset) | |
| 138 | apply (simp add: add_commute) | |
| 139 | done | |
| 140 | ||
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 141 | lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 142 | by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff) | 
| 21164 | 143 | |
| 144 | lemma inverse_diff_inverse: | |
| 145 | "\<lbrakk>(a::'a::division_ring) \<noteq> 0; b \<noteq> 0\<rbrakk> | |
| 146 | \<Longrightarrow> inverse a - inverse b = - (inverse a * (a - b) * inverse b)" | |
| 29667 | 147 | by (simp add: algebra_simps) | 
| 21164 | 148 | |
| 149 | lemma DERIV_inverse_lemma: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 150 | "\<lbrakk>a \<noteq> 0; b \<noteq> (0::'a::real_normed_field)\<rbrakk> | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 151 | \<Longrightarrow> (inverse a - inverse b) / h | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 152 | = - (inverse a * ((a - b) / h) * inverse b)" | 
| 21164 | 153 | by (simp add: inverse_diff_inverse) | 
| 154 | ||
| 155 | lemma DERIV_inverse': | |
| 156 | assumes der: "DERIV f x :> D" | |
| 157 | assumes neq: "f x \<noteq> 0" | |
| 158 | shows "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))" | |
| 159 | (is "DERIV _ _ :> ?E") | |
| 160 | proof (unfold DERIV_iff2) | |
| 161 | from der have lim_f: "f -- x --> f x" | |
| 162 | by (rule DERIV_isCont [unfolded isCont_def]) | |
| 163 | ||
| 164 | from neq have "0 < norm (f x)" by simp | |
| 165 | with LIM_D [OF lim_f] obtain s | |
| 166 | where s: "0 < s" | |
| 167 | and less_fx: "\<And>z. \<lbrakk>z \<noteq> x; norm (z - x) < s\<rbrakk> | |
| 168 | \<Longrightarrow> norm (f z - f x) < norm (f x)" | |
| 169 | by fast | |
| 170 | ||
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 171 | show "(\<lambda>z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E" | 
| 21164 | 172 | proof (rule LIM_equal2 [OF s]) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 173 | fix z | 
| 21164 | 174 | assume "z \<noteq> x" "norm (z - x) < s" | 
| 175 | hence "norm (f z - f x) < norm (f x)" by (rule less_fx) | |
| 176 | hence "f z \<noteq> 0" by auto | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 177 | thus "(inverse (f z) - inverse (f x)) / (z - x) = | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 178 | - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))" | 
| 21164 | 179 | using neq by (rule DERIV_inverse_lemma) | 
| 180 | next | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 181 | from der have "(\<lambda>z. (f z - f x) / (z - x)) -- x --> D" | 
| 21164 | 182 | by (unfold DERIV_iff2) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 183 | thus "(\<lambda>z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))) | 
| 21164 | 184 | -- x --> ?E" | 
| 22613 | 185 | by (intro LIM_mult LIM_inverse LIM_minus LIM_const lim_f neq) | 
| 21164 | 186 | qed | 
| 187 | qed | |
| 188 | ||
| 189 | lemma DERIV_divide: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 190 | "\<lbrakk>DERIV f x :> D; DERIV g x :> E; g x \<noteq> 0\<rbrakk> | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 191 | \<Longrightarrow> DERIV (\<lambda>x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)" | 
| 21164 | 192 | apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) + | 
| 193 | D * inverse (g x) = (D * g x - f x * E) / (g x * g x)") | |
| 194 | apply (erule subst) | |
| 195 | apply (unfold divide_inverse) | |
| 196 | apply (erule DERIV_mult') | |
| 197 | apply (erule (1) DERIV_inverse') | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23441diff
changeset | 198 | apply (simp add: ring_distribs nonzero_inverse_mult_distrib) | 
| 21164 | 199 | apply (simp add: mult_ac) | 
| 200 | done | |
| 201 | ||
| 202 | lemma DERIV_power_Suc: | |
| 31017 | 203 |   fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
 | 
| 21164 | 204 | assumes f: "DERIV f x :> D" | 
| 23431 
25ca91279a9b
change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
 huffman parents: 
23413diff
changeset | 205 | shows "DERIV (\<lambda>x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)" | 
| 21164 | 206 | proof (induct n) | 
| 207 | case 0 | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 208 | show ?case by (simp add: f) | 
| 21164 | 209 | case (Suc k) | 
| 210 | from DERIV_mult' [OF f Suc] show ?case | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23441diff
changeset | 211 | apply (simp only: of_nat_Suc ring_distribs mult_1_left) | 
| 29667 | 212 | apply (simp only: power_Suc algebra_simps) | 
| 21164 | 213 | done | 
| 214 | qed | |
| 215 | ||
| 216 | lemma DERIV_power: | |
| 31017 | 217 |   fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
 | 
| 21164 | 218 | assumes f: "DERIV f x :> D" | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 219 | shows "DERIV (\<lambda>x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 220 | by (cases "n", simp, simp add: DERIV_power_Suc f del: power_Suc) | 
| 21164 | 221 | |
| 29975 | 222 | text {* Caratheodory formulation of derivative at a point *}
 | 
| 21164 | 223 | |
| 224 | lemma CARAT_DERIV: | |
| 225 | "(DERIV f x :> l) = | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 226 | (\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isCont g x & g x = l)" | 
| 21164 | 227 | (is "?lhs = ?rhs") | 
| 228 | proof | |
| 229 | assume der: "DERIV f x :> l" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 230 | show "\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) \<and> isCont g x \<and> g x = l" | 
| 21164 | 231 | proof (intro exI conjI) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 232 | let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" | 
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23412diff
changeset | 233 | show "\<forall>z. f z - f x = ?g z * (z-x)" by simp | 
| 21164 | 234 | show "isCont ?g x" using der | 
| 235 | by (simp add: isCont_iff DERIV_iff diff_minus | |
| 236 | cong: LIM_equal [rule_format]) | |
| 237 | show "?g x = l" by simp | |
| 238 | qed | |
| 239 | next | |
| 240 | assume "?rhs" | |
| 241 | then obtain g where | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 242 | "(\<forall>z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast | 
| 21164 | 243 | thus "(DERIV f x :> l)" | 
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23412diff
changeset | 244 | by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong) | 
| 21164 | 245 | qed | 
| 246 | ||
| 247 | lemma DERIV_chain': | |
| 248 | assumes f: "DERIV f x :> D" | |
| 249 | assumes g: "DERIV g (f x) :> E" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 250 | shows "DERIV (\<lambda>x. g (f x)) x :> E * D" | 
| 21164 | 251 | proof (unfold DERIV_iff2) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 252 | obtain d where d: "\<forall>y. g y - g (f x) = d y * (y - f x)" | 
| 21164 | 253 | and cont_d: "isCont d (f x)" and dfx: "d (f x) = E" | 
| 254 | using CARAT_DERIV [THEN iffD1, OF g] by fast | |
| 255 | from f have "f -- x --> f x" | |
| 256 | by (rule DERIV_isCont [unfolded isCont_def]) | |
| 257 | with cont_d have "(\<lambda>z. d (f z)) -- x --> d (f x)" | |
| 21239 | 258 | by (rule isCont_LIM_compose) | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 259 | hence "(\<lambda>z. d (f z) * ((f z - f x) / (z - x))) | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 260 | -- x --> d (f x) * D" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 261 | by (rule LIM_mult [OF _ f [unfolded DERIV_iff2]]) | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 262 | thus "(\<lambda>z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D" | 
| 35216 | 263 | by (simp add: d dfx) | 
| 21164 | 264 | qed | 
| 265 | ||
| 31899 | 266 | text {*
 | 
| 267 | Let's do the standard proof, though theorem | |
| 268 |  @{text "LIM_mult2"} follows from a NS proof
 | |
| 269 | *} | |
| 21164 | 270 | |
| 271 | lemma DERIV_cmult: | |
| 272 | "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" | |
| 273 | by (drule DERIV_mult' [OF DERIV_const], simp) | |
| 274 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 275 | lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 276 | apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 277 | apply (erule DERIV_cmult) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 278 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 279 | |
| 31899 | 280 | text {* Standard version *}
 | 
| 21164 | 281 | lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" | 
| 35216 | 282 | by (drule (1) DERIV_chain', simp add: o_def mult_commute) | 
| 21164 | 283 | |
| 284 | lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db" | |
| 285 | by (auto dest: DERIV_chain simp add: o_def) | |
| 286 | ||
| 31899 | 287 | text {* Derivative of linear multiplication *}
 | 
| 21164 | 288 | lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c" | 
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23044diff
changeset | 289 | by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp) | 
| 21164 | 290 | |
| 291 | lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" | |
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23044diff
changeset | 292 | apply (cut_tac DERIV_power [OF DERIV_ident]) | 
| 35216 | 293 | apply (simp add: real_of_nat_def) | 
| 21164 | 294 | done | 
| 295 | ||
| 31899 | 296 | text {* Power of @{text "-1"} *}
 | 
| 21164 | 297 | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 298 | lemma DERIV_inverse: | 
| 31017 | 299 |   fixes x :: "'a::{real_normed_field}"
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 300 | shows "x \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 301 | by (drule DERIV_inverse' [OF DERIV_ident]) simp | 
| 21164 | 302 | |
| 31899 | 303 | text {* Derivative of inverse *}
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 304 | lemma DERIV_inverse_fun: | 
| 31017 | 305 |   fixes x :: "'a::{real_normed_field}"
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 306 | shows "[| DERIV f x :> d; f(x) \<noteq> 0 |] | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 307 | ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 308 | by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib) | 
| 21164 | 309 | |
| 31899 | 310 | text {* Derivative of quotient *}
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 311 | lemma DERIV_quotient: | 
| 31017 | 312 |   fixes x :: "'a::{real_normed_field}"
 | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 313 | shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> 0 |] | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 314 | ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" | 
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 315 | by (drule (2) DERIV_divide) (simp add: mult_commute) | 
| 21164 | 316 | |
| 29975 | 317 | lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" | 
| 318 | by auto | |
| 319 | ||
| 31899 | 320 | text {* @{text "DERIV_intros"} *}
 | 
| 321 | ML {*
 | |
| 31902 | 322 | structure Deriv_Intros = Named_Thms | 
| 31899 | 323 | ( | 
| 324 | val name = "DERIV_intros" | |
| 325 | val description = "DERIV introduction rules" | |
| 326 | ) | |
| 327 | *} | |
| 31880 | 328 | |
| 31902 | 329 | setup Deriv_Intros.setup | 
| 31880 | 330 | |
| 331 | lemma DERIV_cong: "\<lbrakk> DERIV f x :> X ; X = Y \<rbrakk> \<Longrightarrow> DERIV f x :> Y" | |
| 332 | by simp | |
| 333 | ||
| 334 | declare | |
| 335 | DERIV_const[THEN DERIV_cong, DERIV_intros] | |
| 336 | DERIV_ident[THEN DERIV_cong, DERIV_intros] | |
| 337 | DERIV_add[THEN DERIV_cong, DERIV_intros] | |
| 338 | DERIV_minus[THEN DERIV_cong, DERIV_intros] | |
| 339 | DERIV_mult[THEN DERIV_cong, DERIV_intros] | |
| 340 | DERIV_diff[THEN DERIV_cong, DERIV_intros] | |
| 341 | DERIV_inverse'[THEN DERIV_cong, DERIV_intros] | |
| 342 | DERIV_divide[THEN DERIV_cong, DERIV_intros] | |
| 343 | DERIV_power[where 'a=real, THEN DERIV_cong, | |
| 344 | unfolded real_of_nat_def[symmetric], DERIV_intros] | |
| 345 | DERIV_setsum[THEN DERIV_cong, DERIV_intros] | |
| 22984 | 346 | |
| 31899 | 347 | |
| 22984 | 348 | subsection {* Differentiability predicate *}
 | 
| 21164 | 349 | |
| 29169 | 350 | definition | 
| 351 | differentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool" | |
| 352 | (infixl "differentiable" 60) where | |
| 353 | "f differentiable x = (\<exists>D. DERIV f x :> D)" | |
| 354 | ||
| 355 | lemma differentiableE [elim?]: | |
| 356 | assumes "f differentiable x" | |
| 357 | obtains df where "DERIV f x :> df" | |
| 358 | using prems unfolding differentiable_def .. | |
| 359 | ||
| 21164 | 360 | lemma differentiableD: "f differentiable x ==> \<exists>D. DERIV f x :> D" | 
| 361 | by (simp add: differentiable_def) | |
| 362 | ||
| 363 | lemma differentiableI: "DERIV f x :> D ==> f differentiable x" | |
| 364 | by (force simp add: differentiable_def) | |
| 365 | ||
| 29169 | 366 | lemma differentiable_ident [simp]: "(\<lambda>x. x) differentiable x" | 
| 367 | by (rule DERIV_ident [THEN differentiableI]) | |
| 368 | ||
| 369 | lemma differentiable_const [simp]: "(\<lambda>z. a) differentiable x" | |
| 370 | by (rule DERIV_const [THEN differentiableI]) | |
| 21164 | 371 | |
| 29169 | 372 | lemma differentiable_compose: | 
| 373 | assumes f: "f differentiable (g x)" | |
| 374 | assumes g: "g differentiable x" | |
| 375 | shows "(\<lambda>x. f (g x)) differentiable x" | |
| 376 | proof - | |
| 377 | from `f differentiable (g x)` obtain df where "DERIV f (g x) :> df" .. | |
| 378 | moreover | |
| 379 | from `g differentiable x` obtain dg where "DERIV g x :> dg" .. | |
| 380 | ultimately | |
| 381 | have "DERIV (\<lambda>x. f (g x)) x :> df * dg" by (rule DERIV_chain2) | |
| 382 | thus ?thesis by (rule differentiableI) | |
| 383 | qed | |
| 384 | ||
| 385 | lemma differentiable_sum [simp]: | |
| 21164 | 386 | assumes "f differentiable x" | 
| 387 | and "g differentiable x" | |
| 388 | shows "(\<lambda>x. f x + g x) differentiable x" | |
| 389 | proof - | |
| 29169 | 390 | from `f differentiable x` obtain df where "DERIV f x :> df" .. | 
| 391 | moreover | |
| 392 | from `g differentiable x` obtain dg where "DERIV g x :> dg" .. | |
| 393 | ultimately | |
| 394 | have "DERIV (\<lambda>x. f x + g x) x :> df + dg" by (rule DERIV_add) | |
| 395 | thus ?thesis by (rule differentiableI) | |
| 396 | qed | |
| 397 | ||
| 398 | lemma differentiable_minus [simp]: | |
| 399 | assumes "f differentiable x" | |
| 400 | shows "(\<lambda>x. - f x) differentiable x" | |
| 401 | proof - | |
| 402 | from `f differentiable x` obtain df where "DERIV f x :> df" .. | |
| 403 | hence "DERIV (\<lambda>x. - f x) x :> - df" by (rule DERIV_minus) | |
| 404 | thus ?thesis by (rule differentiableI) | |
| 21164 | 405 | qed | 
| 406 | ||
| 29169 | 407 | lemma differentiable_diff [simp]: | 
| 21164 | 408 | assumes "f differentiable x" | 
| 29169 | 409 | assumes "g differentiable x" | 
| 21164 | 410 | shows "(\<lambda>x. f x - g x) differentiable x" | 
| 29169 | 411 | unfolding diff_minus using prems by simp | 
| 412 | ||
| 413 | lemma differentiable_mult [simp]: | |
| 414 | assumes "f differentiable x" | |
| 415 | assumes "g differentiable x" | |
| 416 | shows "(\<lambda>x. f x * g x) differentiable x" | |
| 21164 | 417 | proof - | 
| 29169 | 418 | from `f differentiable x` obtain df where "DERIV f x :> df" .. | 
| 21164 | 419 | moreover | 
| 29169 | 420 | from `g differentiable x` obtain dg where "DERIV g x :> dg" .. | 
| 421 | ultimately | |
| 422 | have "DERIV (\<lambda>x. f x * g x) x :> df * g x + dg * f x" by (rule DERIV_mult) | |
| 423 | thus ?thesis by (rule differentiableI) | |
| 21164 | 424 | qed | 
| 425 | ||
| 29169 | 426 | lemma differentiable_inverse [simp]: | 
| 427 | assumes "f differentiable x" and "f x \<noteq> 0" | |
| 428 | shows "(\<lambda>x. inverse (f x)) differentiable x" | |
| 21164 | 429 | proof - | 
| 29169 | 430 | from `f differentiable x` obtain df where "DERIV f x :> df" .. | 
| 431 | hence "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * df * inverse (f x))" | |
| 432 | using `f x \<noteq> 0` by (rule DERIV_inverse') | |
| 433 | thus ?thesis by (rule differentiableI) | |
| 21164 | 434 | qed | 
| 435 | ||
| 29169 | 436 | lemma differentiable_divide [simp]: | 
| 437 | assumes "f differentiable x" | |
| 438 | assumes "g differentiable x" and "g x \<noteq> 0" | |
| 439 | shows "(\<lambda>x. f x / g x) differentiable x" | |
| 440 | unfolding divide_inverse using prems by simp | |
| 441 | ||
| 442 | lemma differentiable_power [simp]: | |
| 31017 | 443 |   fixes f :: "'a::{real_normed_field} \<Rightarrow> 'a"
 | 
| 29169 | 444 | assumes "f differentiable x" | 
| 445 | shows "(\<lambda>x. f x ^ n) differentiable x" | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
30242diff
changeset | 446 | by (induct n, simp, simp add: prems) | 
| 29169 | 447 | |
| 22984 | 448 | |
| 21164 | 449 | subsection {* Nested Intervals and Bisection *}
 | 
| 450 | ||
| 451 | text{*Lemmas about nested intervals and proof by bisection (cf.Harrison).
 | |
| 452 | All considerably tidied by lcp.*} | |
| 453 | ||
| 454 | lemma lemma_f_mono_add [rule_format (no_asm)]: "(\<forall>n. (f::nat=>real) n \<le> f (Suc n)) --> f m \<le> f(m + no)" | |
| 455 | apply (induct "no") | |
| 456 | apply (auto intro: order_trans) | |
| 457 | done | |
| 458 | ||
| 459 | lemma f_inc_g_dec_Beq_f: "[| \<forall>n. f(n) \<le> f(Suc n); | |
| 460 | \<forall>n. g(Suc n) \<le> g(n); | |
| 461 | \<forall>n. f(n) \<le> g(n) |] | |
| 462 | ==> Bseq (f :: nat \<Rightarrow> real)" | |
| 463 | apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI) | |
| 464 | apply (induct_tac "n") | |
| 465 | apply (auto intro: order_trans) | |
| 466 | apply (rule_tac y = "g (Suc na)" in order_trans) | |
| 467 | apply (induct_tac [2] "na") | |
| 468 | apply (auto intro: order_trans) | |
| 469 | done | |
| 470 | ||
| 471 | lemma f_inc_g_dec_Beq_g: "[| \<forall>n. f(n) \<le> f(Suc n); | |
| 472 | \<forall>n. g(Suc n) \<le> g(n); | |
| 473 | \<forall>n. f(n) \<le> g(n) |] | |
| 474 | ==> Bseq (g :: nat \<Rightarrow> real)" | |
| 475 | apply (subst Bseq_minus_iff [symmetric]) | |
| 476 | apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f) | |
| 477 | apply auto | |
| 478 | done | |
| 479 | ||
| 480 | lemma f_inc_imp_le_lim: | |
| 481 | fixes f :: "nat \<Rightarrow> real" | |
| 482 | shows "\<lbrakk>\<forall>n. f n \<le> f (Suc n); convergent f\<rbrakk> \<Longrightarrow> f n \<le> lim f" | |
| 483 | apply (rule linorder_not_less [THEN iffD1]) | |
| 484 | apply (auto simp add: convergent_LIMSEQ_iff LIMSEQ_iff monoseq_Suc) | |
| 485 | apply (drule real_less_sum_gt_zero) | |
| 486 | apply (drule_tac x = "f n + - lim f" in spec, safe) | |
| 487 | apply (drule_tac P = "%na. no\<le>na --> ?Q na" and x = "no + n" in spec, auto) | |
| 488 | apply (subgoal_tac "lim f \<le> f (no + n) ") | |
| 489 | apply (drule_tac no=no and m=n in lemma_f_mono_add) | |
| 490 | apply (auto simp add: add_commute) | |
| 491 | apply (induct_tac "no") | |
| 492 | apply simp | |
| 493 | apply (auto intro: order_trans simp add: diff_minus abs_if) | |
| 494 | done | |
| 495 | ||
| 31404 | 496 | lemma lim_uminus: | 
| 497 | fixes g :: "nat \<Rightarrow> 'a::real_normed_vector" | |
| 498 | shows "convergent g ==> lim (%x. - g x) = - (lim g)" | |
| 21164 | 499 | apply (rule LIMSEQ_minus [THEN limI]) | 
| 500 | apply (simp add: convergent_LIMSEQ_iff) | |
| 501 | done | |
| 502 | ||
| 503 | lemma g_dec_imp_lim_le: | |
| 504 | fixes g :: "nat \<Rightarrow> real" | |
| 505 | shows "\<lbrakk>\<forall>n. g (Suc n) \<le> g(n); convergent g\<rbrakk> \<Longrightarrow> lim g \<le> g n" | |
| 506 | apply (subgoal_tac "- (g n) \<le> - (lim g) ") | |
| 507 | apply (cut_tac [2] f = "%x. - (g x)" in f_inc_imp_le_lim) | |
| 508 | apply (auto simp add: lim_uminus convergent_minus_iff [symmetric]) | |
| 509 | done | |
| 510 | ||
| 511 | lemma lemma_nest: "[| \<forall>n. f(n) \<le> f(Suc n); | |
| 512 | \<forall>n. g(Suc n) \<le> g(n); | |
| 513 | \<forall>n. f(n) \<le> g(n) |] | |
| 514 | ==> \<exists>l m :: real. l \<le> m & ((\<forall>n. f(n) \<le> l) & f ----> l) & | |
| 515 | ((\<forall>n. m \<le> g(n)) & g ----> m)" | |
| 516 | apply (subgoal_tac "monoseq f & monoseq g") | |
| 517 | prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc) | |
| 518 | apply (subgoal_tac "Bseq f & Bseq g") | |
| 519 | prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g) | |
| 520 | apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff) | |
| 521 | apply (rule_tac x = "lim f" in exI) | |
| 522 | apply (rule_tac x = "lim g" in exI) | |
| 523 | apply (auto intro: LIMSEQ_le) | |
| 524 | apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff) | |
| 525 | done | |
| 526 | ||
| 527 | lemma lemma_nest_unique: "[| \<forall>n. f(n) \<le> f(Suc n); | |
| 528 | \<forall>n. g(Suc n) \<le> g(n); | |
| 529 | \<forall>n. f(n) \<le> g(n); | |
| 530 | (%n. f(n) - g(n)) ----> 0 |] | |
| 531 | ==> \<exists>l::real. ((\<forall>n. f(n) \<le> l) & f ----> l) & | |
| 532 | ((\<forall>n. l \<le> g(n)) & g ----> l)" | |
| 533 | apply (drule lemma_nest, auto) | |
| 534 | apply (subgoal_tac "l = m") | |
| 535 | apply (drule_tac [2] X = f in LIMSEQ_diff) | |
| 536 | apply (auto intro: LIMSEQ_unique) | |
| 537 | done | |
| 538 | ||
| 539 | text{*The universal quantifiers below are required for the declaration
 | |
| 540 |   of @{text Bolzano_nest_unique} below.*}
 | |
| 541 | ||
| 542 | lemma Bolzano_bisect_le: | |
| 543 | "a \<le> b ==> \<forall>n. fst (Bolzano_bisect P a b n) \<le> snd (Bolzano_bisect P a b n)" | |
| 544 | apply (rule allI) | |
| 545 | apply (induct_tac "n") | |
| 546 | apply (auto simp add: Let_def split_def) | |
| 547 | done | |
| 548 | ||
| 549 | lemma Bolzano_bisect_fst_le_Suc: "a \<le> b ==> | |
| 550 | \<forall>n. fst(Bolzano_bisect P a b n) \<le> fst (Bolzano_bisect P a b (Suc n))" | |
| 551 | apply (rule allI) | |
| 552 | apply (induct_tac "n") | |
| 553 | apply (auto simp add: Bolzano_bisect_le Let_def split_def) | |
| 554 | done | |
| 555 | ||
| 556 | lemma Bolzano_bisect_Suc_le_snd: "a \<le> b ==> | |
| 557 | \<forall>n. snd(Bolzano_bisect P a b (Suc n)) \<le> snd (Bolzano_bisect P a b n)" | |
| 558 | apply (rule allI) | |
| 559 | apply (induct_tac "n") | |
| 560 | apply (auto simp add: Bolzano_bisect_le Let_def split_def) | |
| 561 | done | |
| 562 | ||
| 563 | lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)" | |
| 564 | apply (auto) | |
| 565 | apply (drule_tac f = "%u. (1/2) *u" in arg_cong) | |
| 566 | apply (simp) | |
| 567 | done | |
| 568 | ||
| 569 | lemma Bolzano_bisect_diff: | |
| 570 | "a \<le> b ==> | |
| 571 | snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = | |
| 572 | (b-a) / (2 ^ n)" | |
| 573 | apply (induct "n") | |
| 574 | apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def) | |
| 575 | done | |
| 576 | ||
| 577 | lemmas Bolzano_nest_unique = | |
| 578 | lemma_nest_unique | |
| 579 | [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le] | |
| 580 | ||
| 581 | ||
| 582 | lemma not_P_Bolzano_bisect: | |
| 583 | assumes P: "!!a b c. [| P(a,b); P(b,c); a \<le> b; b \<le> c|] ==> P(a,c)" | |
| 584 | and notP: "~ P(a,b)" | |
| 585 | and le: "a \<le> b" | |
| 586 | shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" | |
| 587 | proof (induct n) | |
| 23441 | 588 | case 0 show ?case using notP by simp | 
| 21164 | 589 | next | 
| 590 | case (Suc n) | |
| 591 | thus ?case | |
| 592 | by (auto simp del: surjective_pairing [symmetric] | |
| 593 | simp add: Let_def split_def Bolzano_bisect_le [OF le] | |
| 594 | P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"]) | |
| 595 | qed | |
| 596 | ||
| 597 | (*Now we re-package P_prem as a formula*) | |
| 598 | lemma not_P_Bolzano_bisect': | |
| 599 | "[| \<forall>a b c. P(a,b) & P(b,c) & a \<le> b & b \<le> c --> P(a,c); | |
| 600 | ~ P(a,b); a \<le> b |] ==> | |
| 601 | \<forall>n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))" | |
| 602 | by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE]) | |
| 603 | ||
| 604 | ||
| 605 | ||
| 606 | lemma lemma_BOLZANO: | |
| 607 | "[| \<forall>a b c. P(a,b) & P(b,c) & a \<le> b & b \<le> c --> P(a,c); | |
| 608 | \<forall>x. \<exists>d::real. 0 < d & | |
| 609 | (\<forall>a b. a \<le> x & x \<le> b & (b-a) < d --> P(a,b)); | |
| 610 | a \<le> b |] | |
| 611 | ==> P(a,b)" | |
| 612 | apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+) | |
| 613 | apply (rule LIMSEQ_minus_cancel) | |
| 614 | apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero) | |
| 615 | apply (rule ccontr) | |
| 616 | apply (drule not_P_Bolzano_bisect', assumption+) | |
| 617 | apply (rename_tac "l") | |
| 618 | apply (drule_tac x = l in spec, clarify) | |
| 31336 
e17f13cd1280
generalize constants in SEQ.thy to class metric_space
 huffman parents: 
31017diff
changeset | 619 | apply (simp add: LIMSEQ_iff) | 
| 21164 | 620 | apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec) | 
| 621 | apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec) | |
| 622 | apply (drule real_less_half_sum, auto) | |
| 623 | apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec) | |
| 624 | apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec) | |
| 625 | apply safe | |
| 626 | apply (simp_all (no_asm_simp)) | |
| 627 | apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans) | |
| 628 | apply (simp (no_asm_simp) add: abs_if) | |
| 629 | apply (rule real_sum_of_halves [THEN subst]) | |
| 630 | apply (rule add_strict_mono) | |
| 631 | apply (simp_all add: diff_minus [symmetric]) | |
| 632 | done | |
| 633 | ||
| 634 | ||
| 635 | lemma lemma_BOLZANO2: "((\<forall>a b c. (a \<le> b & b \<le> c & P(a,b) & P(b,c)) --> P(a,c)) & | |
| 636 | (\<forall>x. \<exists>d::real. 0 < d & | |
| 637 | (\<forall>a b. a \<le> x & x \<le> b & (b-a) < d --> P(a,b)))) | |
| 638 | --> (\<forall>a b. a \<le> b --> P(a,b))" | |
| 639 | apply clarify | |
| 640 | apply (blast intro: lemma_BOLZANO) | |
| 641 | done | |
| 642 | ||
| 643 | ||
| 644 | subsection {* Intermediate Value Theorem *}
 | |
| 645 | ||
| 646 | text {*Prove Contrapositive by Bisection*}
 | |
| 647 | ||
| 648 | lemma IVT: "[| f(a::real) \<le> (y::real); y \<le> f(b); | |
| 649 | a \<le> b; | |
| 650 | (\<forall>x. a \<le> x & x \<le> b --> isCont f x) |] | |
| 651 | ==> \<exists>x. a \<le> x & x \<le> b & f(x) = y" | |
| 652 | apply (rule contrapos_pp, assumption) | |
| 653 | apply (cut_tac P = "% (u,v) . a \<le> u & u \<le> v & v \<le> b --> ~ (f (u) \<le> y & y \<le> f (v))" in lemma_BOLZANO2) | |
| 654 | apply safe | |
| 655 | apply simp_all | |
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 656 | apply (simp add: isCont_iff LIM_eq) | 
| 21164 | 657 | apply (rule ccontr) | 
| 658 | apply (subgoal_tac "a \<le> x & x \<le> b") | |
| 659 | prefer 2 | |
| 660 | apply simp | |
| 661 | apply (drule_tac P = "%d. 0<d --> ?P d" and x = 1 in spec, arith) | |
| 662 | apply (drule_tac x = x in spec)+ | |
| 663 | apply simp | |
| 664 | apply (drule_tac P = "%r. ?P r --> (\<exists>s>0. ?Q r s) " and x = "\<bar>y - f x\<bar>" in spec) | |
| 665 | apply safe | |
| 666 | apply simp | |
| 667 | apply (drule_tac x = s in spec, clarify) | |
| 668 | apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe) | |
| 669 | apply (drule_tac x = "ba-x" in spec) | |
| 670 | apply (simp_all add: abs_if) | |
| 671 | apply (drule_tac x = "aa-x" in spec) | |
| 672 | apply (case_tac "x \<le> aa", simp_all) | |
| 673 | done | |
| 674 | ||
| 675 | lemma IVT2: "[| f(b::real) \<le> (y::real); y \<le> f(a); | |
| 676 | a \<le> b; | |
| 677 | (\<forall>x. a \<le> x & x \<le> b --> isCont f x) | |
| 678 | |] ==> \<exists>x. a \<le> x & x \<le> b & f(x) = y" | |
| 679 | apply (subgoal_tac "- f a \<le> -y & -y \<le> - f b", clarify) | |
| 680 | apply (drule IVT [where f = "%x. - f x"], assumption) | |
| 681 | apply (auto intro: isCont_minus) | |
| 682 | done | |
| 683 | ||
| 684 | (*HOL style here: object-level formulations*) | |
| 685 | lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b & | |
| 686 | (\<forall>x. a \<le> x & x \<le> b --> isCont f x)) | |
| 687 | --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)" | |
| 688 | apply (blast intro: IVT) | |
| 689 | done | |
| 690 | ||
| 691 | lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b & | |
| 692 | (\<forall>x. a \<le> x & x \<le> b --> isCont f x)) | |
| 693 | --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)" | |
| 694 | apply (blast intro: IVT2) | |
| 695 | done | |
| 696 | ||
| 29975 | 697 | |
| 698 | subsection {* Boundedness of continuous functions *}
 | |
| 699 | ||
| 21164 | 700 | text{*By bisection, function continuous on closed interval is bounded above*}
 | 
| 701 | ||
| 702 | lemma isCont_bounded: | |
| 703 | "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |] | |
| 704 | ==> \<exists>M::real. \<forall>x::real. a \<le> x & x \<le> b --> f(x) \<le> M" | |
| 705 | apply (cut_tac P = "% (u,v) . a \<le> u & u \<le> v & v \<le> b --> (\<exists>M. \<forall>x. u \<le> x & x \<le> v --> f x \<le> M)" in lemma_BOLZANO2) | |
| 706 | apply safe | |
| 707 | apply simp_all | |
| 708 | apply (rename_tac x xa ya M Ma) | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 709 | apply (metis linorder_not_less order_le_less real_le_trans) | 
| 21164 | 710 | apply (case_tac "a \<le> x & x \<le> b") | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 711 | prefer 2 | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 712 | apply (rule_tac x = 1 in exI, force) | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 713 | apply (simp add: LIM_eq isCont_iff) | 
| 21164 | 714 | apply (drule_tac x = x in spec, auto) | 
| 715 | apply (erule_tac V = "\<forall>M. \<exists>x. a \<le> x & x \<le> b & ~ f x \<le> M" in thin_rl) | |
| 716 | apply (drule_tac x = 1 in spec, auto) | |
| 717 | apply (rule_tac x = s in exI, clarify) | |
| 718 | apply (rule_tac x = "\<bar>f x\<bar> + 1" in exI, clarify) | |
| 719 | apply (drule_tac x = "xa-x" in spec) | |
| 720 | apply (auto simp add: abs_ge_self) | |
| 721 | done | |
| 722 | ||
| 723 | text{*Refine the above to existence of least upper bound*}
 | |
| 724 | ||
| 725 | lemma lemma_reals_complete: "((\<exists>x. x \<in> S) & (\<exists>y. isUb UNIV S (y::real))) --> | |
| 726 | (\<exists>t. isLub UNIV S t)" | |
| 727 | by (blast intro: reals_complete) | |
| 728 | ||
| 729 | lemma isCont_has_Ub: "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |] | |
| 730 | ==> \<exists>M::real. (\<forall>x::real. a \<le> x & x \<le> b --> f(x) \<le> M) & | |
| 731 | (\<forall>N. N < M --> (\<exists>x. a \<le> x & x \<le> b & N < f(x)))" | |
| 732 | apply (cut_tac S = "Collect (%y. \<exists>x. a \<le> x & x \<le> b & y = f x)" | |
| 733 | in lemma_reals_complete) | |
| 734 | apply auto | |
| 735 | apply (drule isCont_bounded, assumption) | |
| 736 | apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def) | |
| 737 | apply (rule exI, auto) | |
| 738 | apply (auto dest!: spec simp add: linorder_not_less) | |
| 739 | done | |
| 740 | ||
| 741 | text{*Now show that it attains its upper bound*}
 | |
| 742 | ||
| 743 | lemma isCont_eq_Ub: | |
| 744 | assumes le: "a \<le> b" | |
| 745 | and con: "\<forall>x::real. a \<le> x & x \<le> b --> isCont f x" | |
| 746 | shows "\<exists>M::real. (\<forall>x. a \<le> x & x \<le> b --> f(x) \<le> M) & | |
| 747 | (\<exists>x. a \<le> x & x \<le> b & f(x) = M)" | |
| 748 | proof - | |
| 749 | from isCont_has_Ub [OF le con] | |
| 750 | obtain M where M1: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M" | |
| 751 | and M2: "!!N. N<M ==> \<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x" by blast | |
| 752 | show ?thesis | |
| 753 | proof (intro exI, intro conjI) | |
| 754 | show " \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M" by (rule M1) | |
| 755 | show "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M" | |
| 756 | proof (rule ccontr) | |
| 757 | assume "\<not> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)" | |
| 758 | with M1 have M3: "\<forall>x. a \<le> x & x \<le> b --> f x < M" | |
| 759 | by (fastsimp simp add: linorder_not_le [symmetric]) | |
| 760 | hence "\<forall>x. a \<le> x & x \<le> b --> isCont (%x. inverse (M - f x)) x" | |
| 761 | by (auto simp add: isCont_inverse isCont_diff con) | |
| 762 | from isCont_bounded [OF le this] | |
| 763 | obtain k where k: "!!x. a \<le> x & x \<le> b --> inverse (M - f x) \<le> k" by auto | |
| 764 | have Minv: "!!x. a \<le> x & x \<le> b --> 0 < inverse (M - f (x))" | |
| 29667 | 765 | by (simp add: M3 algebra_simps) | 
| 21164 | 766 | have "!!x. a \<le> x & x \<le> b --> inverse (M - f x) < k+1" using k | 
| 767 | by (auto intro: order_le_less_trans [of _ k]) | |
| 768 | with Minv | |
| 769 | have "!!x. a \<le> x & x \<le> b --> inverse(k+1) < inverse(inverse(M - f x))" | |
| 770 | by (intro strip less_imp_inverse_less, simp_all) | |
| 771 | hence invlt: "!!x. a \<le> x & x \<le> b --> inverse(k+1) < M - f x" | |
| 772 | by simp | |
| 773 | have "M - inverse (k+1) < M" using k [of a] Minv [of a] le | |
| 774 | by (simp, arith) | |
| 775 | from M2 [OF this] | |
| 776 | obtain x where ax: "a \<le> x & x \<le> b & M - inverse(k+1) < f x" .. | |
| 777 | thus False using invlt [of x] by force | |
| 778 | qed | |
| 779 | qed | |
| 780 | qed | |
| 781 | ||
| 782 | ||
| 783 | text{*Same theorem for lower bound*}
 | |
| 784 | ||
| 785 | lemma isCont_eq_Lb: "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |] | |
| 786 | ==> \<exists>M::real. (\<forall>x::real. a \<le> x & x \<le> b --> M \<le> f(x)) & | |
| 787 | (\<exists>x. a \<le> x & x \<le> b & f(x) = M)" | |
| 788 | apply (subgoal_tac "\<forall>x. a \<le> x & x \<le> b --> isCont (%x. - (f x)) x") | |
| 789 | prefer 2 apply (blast intro: isCont_minus) | |
| 790 | apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub) | |
| 791 | apply safe | |
| 792 | apply auto | |
| 793 | done | |
| 794 | ||
| 795 | ||
| 796 | text{*Another version.*}
 | |
| 797 | ||
| 798 | lemma isCont_Lb_Ub: "[|a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |] | |
| 799 | ==> \<exists>L M::real. (\<forall>x::real. a \<le> x & x \<le> b --> L \<le> f(x) & f(x) \<le> M) & | |
| 800 | (\<forall>y. L \<le> y & y \<le> M --> (\<exists>x. a \<le> x & x \<le> b & (f(x) = y)))" | |
| 801 | apply (frule isCont_eq_Lb) | |
| 802 | apply (frule_tac [2] isCont_eq_Ub) | |
| 803 | apply (assumption+, safe) | |
| 804 | apply (rule_tac x = "f x" in exI) | |
| 805 | apply (rule_tac x = "f xa" in exI, simp, safe) | |
| 806 | apply (cut_tac x = x and y = xa in linorder_linear, safe) | |
| 807 | apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) | |
| 808 | apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) | |
| 809 | apply (rule_tac [2] x = xb in exI) | |
| 810 | apply (rule_tac [4] x = xb in exI, simp_all) | |
| 811 | done | |
| 812 | ||
| 813 | ||
| 29975 | 814 | subsection {* Local extrema *}
 | 
| 815 | ||
| 21164 | 816 | text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
 | 
| 817 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 818 | lemma DERIV_pos_inc_right: | 
| 21164 | 819 | fixes f :: "real => real" | 
| 820 | assumes der: "DERIV f x :> l" | |
| 821 | and l: "0 < l" | |
| 822 | shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x + h)" | |
| 823 | proof - | |
| 824 | from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]] | |
| 825 | have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)" | |
| 826 | by (simp add: diff_minus) | |
| 827 | then obtain s | |
| 828 | where s: "0 < s" | |
| 829 | and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l" | |
| 830 | by auto | |
| 831 | thus ?thesis | |
| 832 | proof (intro exI conjI strip) | |
| 23441 | 833 | show "0<s" using s . | 
| 21164 | 834 | fix h::real | 
| 835 | assume "0 < h" "h < s" | |
| 836 | with all [of h] show "f x < f (x+h)" | |
| 837 | proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric] | |
| 838 | split add: split_if_asm) | |
| 839 | assume "~ (f (x+h) - f x) / h < l" and h: "0 < h" | |
| 840 | with l | |
| 841 | have "0 < (f (x+h) - f x) / h" by arith | |
| 842 | thus "f x < f (x+h)" | |
| 843 | by (simp add: pos_less_divide_eq h) | |
| 844 | qed | |
| 845 | qed | |
| 846 | qed | |
| 847 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 848 | lemma DERIV_neg_dec_left: | 
| 21164 | 849 | fixes f :: "real => real" | 
| 850 | assumes der: "DERIV f x :> l" | |
| 851 | and l: "l < 0" | |
| 852 | shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x-h)" | |
| 853 | proof - | |
| 854 | from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]] | |
| 855 | have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)" | |
| 856 | by (simp add: diff_minus) | |
| 857 | then obtain s | |
| 858 | where s: "0 < s" | |
| 859 | and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l" | |
| 860 | by auto | |
| 861 | thus ?thesis | |
| 862 | proof (intro exI conjI strip) | |
| 23441 | 863 | show "0<s" using s . | 
| 21164 | 864 | fix h::real | 
| 865 | assume "0 < h" "h < s" | |
| 866 | with all [of "-h"] show "f x < f (x-h)" | |
| 867 | proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric] | |
| 868 | split add: split_if_asm) | |
| 869 | assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h" | |
| 870 | with l | |
| 871 | have "0 < (f (x-h) - f x) / h" by arith | |
| 872 | thus "f x < f (x-h)" | |
| 873 | by (simp add: pos_less_divide_eq h) | |
| 874 | qed | |
| 875 | qed | |
| 876 | qed | |
| 877 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 878 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 879 | lemma DERIV_pos_inc_left: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 880 | fixes f :: "real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 881 | shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 882 | apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified]) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 883 | apply (auto simp add: DERIV_minus) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 884 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 885 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 886 | lemma DERIV_neg_dec_right: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 887 | fixes f :: "real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 888 | shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x) > f(x + h)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 889 | apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified]) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 890 | apply (auto simp add: DERIV_minus) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 891 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 892 | |
| 21164 | 893 | lemma DERIV_local_max: | 
| 894 | fixes f :: "real => real" | |
| 895 | assumes der: "DERIV f x :> l" | |
| 896 | and d: "0 < d" | |
| 897 | and le: "\<forall>y. \<bar>x-y\<bar> < d --> f(y) \<le> f(x)" | |
| 898 | shows "l = 0" | |
| 899 | proof (cases rule: linorder_cases [of l 0]) | |
| 23441 | 900 | case equal thus ?thesis . | 
| 21164 | 901 | next | 
| 902 | case less | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 903 | from DERIV_neg_dec_left [OF der less] | 
| 21164 | 904 | obtain d' where d': "0 < d'" | 
| 905 | and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x-h)" by blast | |
| 906 | from real_lbound_gt_zero [OF d d'] | |
| 907 | obtain e where "0 < e \<and> e < d \<and> e < d'" .. | |
| 908 | with lt le [THEN spec [where x="x-e"]] | |
| 909 | show ?thesis by (auto simp add: abs_if) | |
| 910 | next | |
| 911 | case greater | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 912 | from DERIV_pos_inc_right [OF der greater] | 
| 21164 | 913 | obtain d' where d': "0 < d'" | 
| 914 | and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" by blast | |
| 915 | from real_lbound_gt_zero [OF d d'] | |
| 916 | obtain e where "0 < e \<and> e < d \<and> e < d'" .. | |
| 917 | with lt le [THEN spec [where x="x+e"]] | |
| 918 | show ?thesis by (auto simp add: abs_if) | |
| 919 | qed | |
| 920 | ||
| 921 | ||
| 922 | text{*Similar theorem for a local minimum*}
 | |
| 923 | lemma DERIV_local_min: | |
| 924 | fixes f :: "real => real" | |
| 925 | shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0" | |
| 926 | by (drule DERIV_minus [THEN DERIV_local_max], auto) | |
| 927 | ||
| 928 | ||
| 929 | text{*In particular, if a function is locally flat*}
 | |
| 930 | lemma DERIV_local_const: | |
| 931 | fixes f :: "real => real" | |
| 932 | shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0" | |
| 933 | by (auto dest!: DERIV_local_max) | |
| 934 | ||
| 29975 | 935 | |
| 936 | subsection {* Rolle's Theorem *}
 | |
| 937 | ||
| 21164 | 938 | text{*Lemma about introducing open ball in open interval*}
 | 
| 939 | lemma lemma_interval_lt: | |
| 940 | "[| a < x; x < b |] | |
| 941 | ==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)" | |
| 27668 | 942 | |
| 22998 | 943 | apply (simp add: abs_less_iff) | 
| 21164 | 944 | apply (insert linorder_linear [of "x-a" "b-x"], safe) | 
| 945 | apply (rule_tac x = "x-a" in exI) | |
| 946 | apply (rule_tac [2] x = "b-x" in exI, auto) | |
| 947 | done | |
| 948 | ||
| 949 | lemma lemma_interval: "[| a < x; x < b |] ==> | |
| 950 | \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a \<le> y & y \<le> b)" | |
| 951 | apply (drule lemma_interval_lt, auto) | |
| 952 | apply (auto intro!: exI) | |
| 953 | done | |
| 954 | ||
| 955 | text{*Rolle's Theorem.
 | |
| 956 |    If @{term f} is defined and continuous on the closed interval
 | |
| 957 |    @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
 | |
| 958 |    and @{term "f(a) = f(b)"},
 | |
| 959 |    then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
 | |
| 960 | theorem Rolle: | |
| 961 | assumes lt: "a < b" | |
| 962 | and eq: "f(a) = f(b)" | |
| 963 | and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x" | |
| 964 | and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 965 | shows "\<exists>z::real. a < z & z < b & DERIV f z :> 0" | 
| 21164 | 966 | proof - | 
| 967 | have le: "a \<le> b" using lt by simp | |
| 968 | from isCont_eq_Ub [OF le con] | |
| 969 | obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x" | |
| 970 | and alex: "a \<le> x" and xleb: "x \<le> b" | |
| 971 | by blast | |
| 972 | from isCont_eq_Lb [OF le con] | |
| 973 | obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z" | |
| 974 | and alex': "a \<le> x'" and x'leb: "x' \<le> b" | |
| 975 | by blast | |
| 976 | show ?thesis | |
| 977 | proof cases | |
| 978 | assume axb: "a < x & x < b" | |
| 979 |         --{*@{term f} attains its maximum within the interval*}
 | |
| 27668 | 980 | hence ax: "a<x" and xb: "x<b" by arith + | 
| 21164 | 981 | from lemma_interval [OF ax xb] | 
| 982 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 983 | by blast | |
| 984 | hence bound': "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> f y \<le> f x" using x_max | |
| 985 | by blast | |
| 986 | from differentiableD [OF dif [OF axb]] | |
| 987 | obtain l where der: "DERIV f x :> l" .. | |
| 988 | have "l=0" by (rule DERIV_local_max [OF der d bound']) | |
| 989 |         --{*the derivative at a local maximum is zero*}
 | |
| 990 | thus ?thesis using ax xb der by auto | |
| 991 | next | |
| 992 | assume notaxb: "~ (a < x & x < b)" | |
| 993 | hence xeqab: "x=a | x=b" using alex xleb by arith | |
| 994 | hence fb_eq_fx: "f b = f x" by (auto simp add: eq) | |
| 995 | show ?thesis | |
| 996 | proof cases | |
| 997 | assume ax'b: "a < x' & x' < b" | |
| 998 |         --{*@{term f} attains its minimum within the interval*}
 | |
| 27668 | 999 | hence ax': "a<x'" and x'b: "x'<b" by arith+ | 
| 21164 | 1000 | from lemma_interval [OF ax' x'b] | 
| 1001 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 1002 | by blast | |
| 1003 | hence bound': "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> f x' \<le> f y" using x'_min | |
| 1004 | by blast | |
| 1005 | from differentiableD [OF dif [OF ax'b]] | |
| 1006 | obtain l where der: "DERIV f x' :> l" .. | |
| 1007 | have "l=0" by (rule DERIV_local_min [OF der d bound']) | |
| 1008 |         --{*the derivative at a local minimum is zero*}
 | |
| 1009 | thus ?thesis using ax' x'b der by auto | |
| 1010 | next | |
| 1011 | assume notax'b: "~ (a < x' & x' < b)" | |
| 1012 |         --{*@{term f} is constant througout the interval*}
 | |
| 1013 | hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith | |
| 1014 | hence fb_eq_fx': "f b = f x'" by (auto simp add: eq) | |
| 1015 | from dense [OF lt] | |
| 1016 | obtain r where ar: "a < r" and rb: "r < b" by blast | |
| 1017 | from lemma_interval [OF ar rb] | |
| 1018 | obtain d where d: "0<d" and bound: "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b" | |
| 1019 | by blast | |
| 1020 | have eq_fb: "\<forall>z. a \<le> z --> z \<le> b --> f z = f b" | |
| 1021 | proof (clarify) | |
| 1022 | fix z::real | |
| 1023 | assume az: "a \<le> z" and zb: "z \<le> b" | |
| 1024 | show "f z = f b" | |
| 1025 | proof (rule order_antisym) | |
| 1026 | show "f z \<le> f b" by (simp add: fb_eq_fx x_max az zb) | |
| 1027 | show "f b \<le> f z" by (simp add: fb_eq_fx' x'_min az zb) | |
| 1028 | qed | |
| 1029 | qed | |
| 1030 | have bound': "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> f r = f y" | |
| 1031 | proof (intro strip) | |
| 1032 | fix y::real | |
| 1033 | assume lt: "\<bar>r-y\<bar> < d" | |
| 1034 | hence "f y = f b" by (simp add: eq_fb bound) | |
| 1035 | thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le) | |
| 1036 | qed | |
| 1037 | from differentiableD [OF dif [OF conjI [OF ar rb]]] | |
| 1038 | obtain l where der: "DERIV f r :> l" .. | |
| 1039 | have "l=0" by (rule DERIV_local_const [OF der d bound']) | |
| 1040 |         --{*the derivative of a constant function is zero*}
 | |
| 1041 | thus ?thesis using ar rb der by auto | |
| 1042 | qed | |
| 1043 | qed | |
| 1044 | qed | |
| 1045 | ||
| 1046 | ||
| 1047 | subsection{*Mean Value Theorem*}
 | |
| 1048 | ||
| 1049 | lemma lemma_MVT: | |
| 1050 | "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)" | |
| 1051 | proof cases | |
| 1052 | assume "a=b" thus ?thesis by simp | |
| 1053 | next | |
| 1054 | assume "a\<noteq>b" | |
| 1055 | hence ba: "b-a \<noteq> 0" by arith | |
| 1056 | show ?thesis | |
| 1057 | by (rule real_mult_left_cancel [OF ba, THEN iffD1], | |
| 1058 | simp add: right_diff_distrib, | |
| 1059 | simp add: left_diff_distrib) | |
| 1060 | qed | |
| 1061 | ||
| 1062 | theorem MVT: | |
| 1063 | assumes lt: "a < b" | |
| 1064 | and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x" | |
| 1065 | and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x" | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1066 | shows "\<exists>l z::real. a < z & z < b & DERIV f z :> l & | 
| 21164 | 1067 | (f(b) - f(a) = (b-a) * l)" | 
| 1068 | proof - | |
| 1069 | let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" | |
| 1070 | have contF: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?F x" using con | |
| 23069 
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
 huffman parents: 
23044diff
changeset | 1071 | by (fast intro: isCont_diff isCont_const isCont_mult isCont_ident) | 
| 21164 | 1072 | have difF: "\<forall>x. a < x \<and> x < b \<longrightarrow> ?F differentiable x" | 
| 1073 | proof (clarify) | |
| 1074 | fix x::real | |
| 1075 | assume ax: "a < x" and xb: "x < b" | |
| 1076 | from differentiableD [OF dif [OF conjI [OF ax xb]]] | |
| 1077 | obtain l where der: "DERIV f x :> l" .. | |
| 1078 | show "?F differentiable x" | |
| 1079 | by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"], | |
| 1080 | blast intro: DERIV_diff DERIV_cmult_Id der) | |
| 1081 | qed | |
| 1082 | from Rolle [where f = ?F, OF lt lemma_MVT contF difF] | |
| 1083 | obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0" | |
| 1084 | by blast | |
| 1085 | have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)" | |
| 1086 | by (rule DERIV_cmult_Id) | |
| 1087 | hence derF: "DERIV (\<lambda>x. ?F x + (f b - f a) / (b - a) * x) z | |
| 1088 | :> 0 + (f b - f a) / (b - a)" | |
| 1089 | by (rule DERIV_add [OF der]) | |
| 1090 | show ?thesis | |
| 1091 | proof (intro exI conjI) | |
| 23441 | 1092 | show "a < z" using az . | 
| 1093 | show "z < b" using zb . | |
| 21164 | 1094 | show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp) | 
| 1095 | show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp | |
| 1096 | qed | |
| 1097 | qed | |
| 1098 | ||
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1099 | lemma MVT2: | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1100 | "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |] | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1101 | ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1102 | apply (drule MVT) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1103 | apply (blast intro: DERIV_isCont) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1104 | apply (force dest: order_less_imp_le simp add: differentiable_def) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1105 | apply (blast dest: DERIV_unique order_less_imp_le) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1106 | done | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1107 | |
| 21164 | 1108 | |
| 1109 | text{*A function is constant if its derivative is 0 over an interval.*}
 | |
| 1110 | ||
| 1111 | lemma DERIV_isconst_end: | |
| 1112 | fixes f :: "real => real" | |
| 1113 | shows "[| a < b; | |
| 1114 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1115 | \<forall>x. a < x & x < b --> DERIV f x :> 0 |] | |
| 1116 | ==> f b = f a" | |
| 1117 | apply (drule MVT, assumption) | |
| 1118 | apply (blast intro: differentiableI) | |
| 1119 | apply (auto dest!: DERIV_unique simp add: diff_eq_eq) | |
| 1120 | done | |
| 1121 | ||
| 1122 | lemma DERIV_isconst1: | |
| 1123 | fixes f :: "real => real" | |
| 1124 | shows "[| a < b; | |
| 1125 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1126 | \<forall>x. a < x & x < b --> DERIV f x :> 0 |] | |
| 1127 | ==> \<forall>x. a \<le> x & x \<le> b --> f x = f a" | |
| 1128 | apply safe | |
| 1129 | apply (drule_tac x = a in order_le_imp_less_or_eq, safe) | |
| 1130 | apply (drule_tac b = x in DERIV_isconst_end, auto) | |
| 1131 | done | |
| 1132 | ||
| 1133 | lemma DERIV_isconst2: | |
| 1134 | fixes f :: "real => real" | |
| 1135 | shows "[| a < b; | |
| 1136 | \<forall>x. a \<le> x & x \<le> b --> isCont f x; | |
| 1137 | \<forall>x. a < x & x < b --> DERIV f x :> 0; | |
| 1138 | a \<le> x; x \<le> b |] | |
| 1139 | ==> f x = f a" | |
| 1140 | apply (blast dest: DERIV_isconst1) | |
| 1141 | done | |
| 1142 | ||
| 29803 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1143 | lemma DERIV_isconst3: fixes a b x y :: real | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1144 |   assumes "a < b" and "x \<in> {a <..< b}" and "y \<in> {a <..< b}"
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1145 |   assumes derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1146 | 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 | 1147 | 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 | 1148 | case False | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1149 | 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 | 1150 | let ?b = "max x y" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1151 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1152 | have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1153 | proof (rule allI, rule impI) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1154 | fix z :: real assume "?a \<le> z \<and> z \<le> ?b" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1155 |     hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1156 |     hence "z \<in> {a<..<b}" by auto
 | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1157 | thus "DERIV f z :> 0" by (rule derivable) | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1158 | qed | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1159 | hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z" | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1160 | and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1161 | |
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1162 | have "?a < ?b" using `x \<noteq> y` by auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1163 | 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 | 1164 | 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 | 1165 | qed auto | 
| 
c56a5571f60a
Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
 hoelzl parents: 
29667diff
changeset | 1166 | |
| 21164 | 1167 | lemma DERIV_isconst_all: | 
| 1168 | fixes f :: "real => real" | |
| 1169 | shows "\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)" | |
| 1170 | apply (rule linorder_cases [of x y]) | |
| 1171 | apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+ | |
| 1172 | done | |
| 1173 | ||
| 1174 | lemma DERIV_const_ratio_const: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1175 | fixes f :: "real => real" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1176 | shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k" | 
| 21164 | 1177 | apply (rule linorder_cases [of a b], auto) | 
| 1178 | apply (drule_tac [!] f = f in MVT) | |
| 1179 | apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23441diff
changeset | 1180 | apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus) | 
| 21164 | 1181 | done | 
| 1182 | ||
| 1183 | lemma DERIV_const_ratio_const2: | |
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1184 | fixes f :: "real => real" | 
| 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1185 | shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k" | 
| 21164 | 1186 | apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1]) | 
| 1187 | apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc) | |
| 1188 | done | |
| 1189 | ||
| 1190 | lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)" | |
| 1191 | by (simp) | |
| 1192 | ||
| 1193 | lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)" | |
| 1194 | by (simp) | |
| 1195 | ||
| 1196 | text{*Gallileo's "trick": average velocity = av. of end velocities*}
 | |
| 1197 | ||
| 1198 | lemma DERIV_const_average: | |
| 1199 | fixes v :: "real => real" | |
| 1200 | assumes neq: "a \<noteq> (b::real)" | |
| 1201 | and der: "\<forall>x. DERIV v x :> k" | |
| 1202 | shows "v ((a + b)/2) = (v a + v b)/2" | |
| 1203 | proof (cases rule: linorder_cases [of a b]) | |
| 1204 | case equal with neq show ?thesis by simp | |
| 1205 | next | |
| 1206 | case less | |
| 1207 | have "(v b - v a) / (b - a) = k" | |
| 1208 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 1209 | hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp | |
| 1210 | moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" | |
| 1211 | by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) | |
| 1212 | ultimately show ?thesis using neq by force | |
| 1213 | next | |
| 1214 | case greater | |
| 1215 | have "(v b - v a) / (b - a) = k" | |
| 1216 | by (rule DERIV_const_ratio_const2 [OF neq der]) | |
| 1217 | hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp | |
| 1218 | moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" | |
| 1219 | by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq) | |
| 1220 | ultimately show ?thesis using neq by (force simp add: add_commute) | |
| 1221 | qed | |
| 1222 | ||
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1223 | (* A function with positive derivative is increasing. | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1224 | A simple proof using the MVT, by Jeremy Avigad. And variants. | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1225 | *) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1226 | lemma DERIV_pos_imp_increasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1227 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1228 | assumes "a < b" and "\<forall>x. a \<le> x & x \<le> b --> (EX y. DERIV f x :> y & y > 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1229 | shows "f a < f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1230 | proof (rule ccontr) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1231 | assume "~ f a < f b" | 
| 33690 | 1232 | have "EX l z. a < z & z < b & DERIV f z :> l | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1233 | & f b - f a = (b - a) * l" | 
| 33690 | 1234 | apply (rule MVT) | 
| 1235 | using assms | |
| 1236 | apply auto | |
| 1237 | apply (metis DERIV_isCont) | |
| 1238 | apply (metis differentiableI real_less_def) | |
| 1239 | done | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1240 | then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1241 | and "f b - f a = (b - a) * l" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1242 | by auto | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1243 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1244 | from prems have "~(l > 0)" | 
| 33690 | 1245 | by (metis linorder_not_le mult_le_0_iff real_le_eq_diff) | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1246 | with prems show False | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1247 | by (metis DERIV_unique real_less_def) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1248 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1249 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1250 | lemma DERIV_nonneg_imp_nonincreasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1251 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1252 | assumes "a \<le> b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1253 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<ge> 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1254 | shows "f a \<le> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1255 | proof (rule ccontr, cases "a = b") | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1256 | assume "~ f a \<le> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1257 | assume "a = b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1258 | with prems show False by auto | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1259 | next assume "~ f a \<le> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1260 | assume "a ~= b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1261 | with assms have "EX l z. a < z & z < b & DERIV f z :> l | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1262 | & f b - f a = (b - a) * l" | 
| 33690 | 1263 | apply - | 
| 1264 | apply (rule MVT) | |
| 1265 | apply auto | |
| 1266 | apply (metis DERIV_isCont) | |
| 1267 | apply (metis differentiableI real_less_def) | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1268 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1269 | then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1270 | and "f b - f a = (b - a) * l" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1271 | by auto | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1272 | from prems have "~(l >= 0)" | 
| 33659 | 1273 | by (metis diff_self le_eqI le_iff_diff_le_0 real_le_antisym real_le_linear | 
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1274 | split_mult_pos_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1275 | with prems show False | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1276 | by (metis DERIV_unique order_less_imp_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1277 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1278 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1279 | lemma DERIV_neg_imp_decreasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1280 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1281 | assumes "a < b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1282 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y < 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1283 | shows "f a > f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1284 | proof - | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1285 | have "(%x. -f x) a < (%x. -f x) b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1286 | apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"]) | 
| 33690 | 1287 | using assms | 
| 1288 | apply auto | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1289 | apply (metis DERIV_minus neg_0_less_iff_less) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1290 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1291 | thus ?thesis | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1292 | by simp | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1293 | qed | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1294 | |
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1295 | lemma DERIV_nonpos_imp_nonincreasing: | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1296 | fixes a::real and b::real and f::"real => real" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1297 | assumes "a \<le> b" and | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1298 | "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<le> 0)" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1299 | shows "f a \<ge> f b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1300 | proof - | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1301 | have "(%x. -f x) a \<le> (%x. -f x) b" | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1302 | apply (rule DERIV_nonneg_imp_nonincreasing [of a b "%x. -f x"]) | 
| 33690 | 1303 | using assms | 
| 1304 | apply auto | |
| 33654 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1305 | apply (metis DERIV_minus neg_0_le_iff_le) | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1306 | done | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1307 | thus ?thesis | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1308 | by simp | 
| 
abf780db30ea
A number of theorems contributed by Jeremy Avigad
 paulson parents: 
31902diff
changeset | 1309 | qed | 
| 21164 | 1310 | |
| 29975 | 1311 | subsection {* Continuous injective functions *}
 | 
| 1312 | ||
| 21164 | 1313 | text{*Dull lemma: an continuous injection on an interval must have a
 | 
| 1314 | strict maximum at an end point, not in the middle.*} | |
| 1315 | ||
| 1316 | lemma lemma_isCont_inj: | |
| 1317 | fixes f :: "real \<Rightarrow> real" | |
| 1318 | assumes d: "0 < d" | |
| 1319 | and inj [rule_format]: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z" | |
| 1320 | and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z" | |
| 1321 | shows "\<exists>z. \<bar>z-x\<bar> \<le> d & f x < f z" | |
| 1322 | proof (rule ccontr) | |
| 1323 | assume "~ (\<exists>z. \<bar>z-x\<bar> \<le> d & f x < f z)" | |
| 1324 | hence all [rule_format]: "\<forall>z. \<bar>z - x\<bar> \<le> d --> f z \<le> f x" by auto | |
| 1325 | show False | |
| 1326 | proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"]) | |
| 1327 | case le | |
| 1328 | from d cont all [of "x+d"] | |
| 1329 | have flef: "f(x+d) \<le> f x" | |
| 1330 | and xlex: "x - d \<le> x" | |
| 1331 | and cont': "\<forall>z. x - d \<le> z \<and> z \<le> x \<longrightarrow> isCont f z" | |
| 1332 | by (auto simp add: abs_if) | |
| 1333 | from IVT [OF le flef xlex cont'] | |
| 1334 | obtain x' where "x-d \<le> x'" "x' \<le> x" "f x' = f(x+d)" by blast | |
| 1335 | moreover | |
| 1336 | hence "g(f x') = g (f(x+d))" by simp | |
| 1337 | ultimately show False using d inj [of x'] inj [of "x+d"] | |
| 22998 | 1338 | by (simp add: abs_le_iff) | 
| 21164 | 1339 | next | 
| 1340 | case ge | |
| 1341 | from d cont all [of "x-d"] | |
| 1342 | have flef: "f(x-d) \<le> f x" | |
| 1343 | and xlex: "x \<le> x+d" | |
| 1344 | and cont': "\<forall>z. x \<le> z \<and> z \<le> x+d \<longrightarrow> isCont f z" | |
| 1345 | by (auto simp add: abs_if) | |
| 1346 | from IVT2 [OF ge flef xlex cont'] | |
| 1347 | obtain x' where "x \<le> x'" "x' \<le> x+d" "f x' = f(x-d)" by blast | |
| 1348 | moreover | |
| 1349 | hence "g(f x') = g (f(x-d))" by simp | |
| 1350 | ultimately show False using d inj [of x'] inj [of "x-d"] | |
| 22998 | 1351 | by (simp add: abs_le_iff) | 
| 21164 | 1352 | qed | 
| 1353 | qed | |
| 1354 | ||
| 1355 | ||
| 1356 | text{*Similar version for lower bound.*}
 | |
| 1357 | ||
| 1358 | lemma lemma_isCont_inj2: | |
| 1359 | fixes f g :: "real \<Rightarrow> real" | |
| 1360 | shows "[|0 < d; \<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z; | |
| 1361 | \<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z |] | |
| 1362 | ==> \<exists>z. \<bar>z-x\<bar> \<le> d & f z < f x" | |
| 1363 | apply (insert lemma_isCont_inj | |
| 1364 | [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) | |
| 1365 | apply (simp add: isCont_minus linorder_not_le) | |
| 1366 | done | |
| 1367 | ||
| 1368 | text{*Show there's an interval surrounding @{term "f(x)"} in
 | |
| 1369 | @{text "f[[x - d, x + d]]"} .*}
 | |
| 1370 | ||
| 1371 | lemma isCont_inj_range: | |
| 1372 | fixes f :: "real \<Rightarrow> real" | |
| 1373 | assumes d: "0 < d" | |
| 1374 | and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z" | |
| 1375 | and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z" | |
| 1376 | shows "\<exists>e>0. \<forall>y. \<bar>y - f x\<bar> \<le> e --> (\<exists>z. \<bar>z-x\<bar> \<le> d & f z = y)" | |
| 1377 | proof - | |
| 1378 | have "x-d \<le> x+d" "\<forall>z. x-d \<le> z \<and> z \<le> x+d \<longrightarrow> isCont f z" using cont d | |
| 22998 | 1379 | by (auto simp add: abs_le_iff) | 
| 21164 | 1380 | from isCont_Lb_Ub [OF this] | 
| 1381 | obtain L M | |
| 1382 | where all1 [rule_format]: "\<forall>z. x-d \<le> z \<and> z \<le> x+d \<longrightarrow> L \<le> f z \<and> f z \<le> M" | |
| 1383 | and all2 [rule_format]: | |
| 1384 | "\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>z. x-d \<le> z \<and> z \<le> x+d \<and> f z = y)" | |
| 1385 | by auto | |
| 1386 | with d have "L \<le> f x & f x \<le> M" by simp | |
| 1387 | moreover have "L \<noteq> f x" | |
| 1388 | proof - | |
| 1389 | from lemma_isCont_inj2 [OF d inj cont] | |
| 1390 | obtain u where "\<bar>u - x\<bar> \<le> d" "f u < f x" by auto | |
| 1391 | thus ?thesis using all1 [of u] by arith | |
| 1392 | qed | |
| 1393 | moreover have "f x \<noteq> M" | |
| 1394 | proof - | |
| 1395 | from lemma_isCont_inj [OF d inj cont] | |
| 1396 | obtain u where "\<bar>u - x\<bar> \<le> d" "f x < f u" by auto | |
| 1397 | thus ?thesis using all1 [of u] by arith | |
| 1398 | qed | |
| 1399 | ultimately have "L < f x & f x < M" by arith | |
| 1400 | hence "0 < f x - L" "0 < M - f x" by arith+ | |
| 1401 | from real_lbound_gt_zero [OF this] | |
| 1402 | obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto | |
| 1403 | thus ?thesis | |
| 1404 | proof (intro exI conjI) | |
| 23441 | 1405 | show "0<e" using e(1) . | 
| 21164 | 1406 | show "\<forall>y. \<bar>y - f x\<bar> \<le> e \<longrightarrow> (\<exists>z. \<bar>z - x\<bar> \<le> d \<and> f z = y)" | 
| 1407 | proof (intro strip) | |
| 1408 | fix y::real | |
| 1409 | assume "\<bar>y - f x\<bar> \<le> e" | |
| 1410 | with e have "L \<le> y \<and> y \<le> M" by arith | |
| 1411 | from all2 [OF this] | |
| 1412 | obtain z where "x - d \<le> z" "z \<le> x + d" "f z = y" by blast | |
| 27668 | 1413 | thus "\<exists>z. \<bar>z - x\<bar> \<le> d \<and> f z = y" | 
| 22998 | 1414 | by (force simp add: abs_le_iff) | 
| 21164 | 1415 | qed | 
| 1416 | qed | |
| 1417 | qed | |
| 1418 | ||
| 1419 | ||
| 1420 | text{*Continuity of inverse function*}
 | |
| 1421 | ||
| 1422 | lemma isCont_inverse_function: | |
| 1423 | fixes f g :: "real \<Rightarrow> real" | |
| 1424 | assumes d: "0 < d" | |
| 1425 | and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z" | |
| 1426 | and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z" | |
| 1427 | shows "isCont g (f x)" | |
| 1428 | proof (simp add: isCont_iff LIM_eq) | |
| 1429 | show "\<forall>r. 0 < r \<longrightarrow> | |
| 1430 | (\<exists>s>0. \<forall>z. z\<noteq>0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>g(f x + z) - g(f x)\<bar> < r)" | |
| 1431 | proof (intro strip) | |
| 1432 | fix r::real | |
| 1433 | assume r: "0<r" | |
| 1434 | from real_lbound_gt_zero [OF r d] | |
| 1435 | obtain e where e: "0 < e" and e_lt: "e < r \<and> e < d" by blast | |
| 1436 | with inj cont | |
| 1437 | have e_simps: "\<forall>z. \<bar>z-x\<bar> \<le> e --> g (f z) = z" | |
| 1438 | "\<forall>z. \<bar>z-x\<bar> \<le> e --> isCont f z" by auto | |
| 1439 | from isCont_inj_range [OF e this] | |
| 1440 | obtain e' where e': "0 < e'" | |
| 1441 | and all: "\<forall>y. \<bar>y - f x\<bar> \<le> e' \<longrightarrow> (\<exists>z. \<bar>z - x\<bar> \<le> e \<and> f z = y)" | |
| 1442 | by blast | |
| 1443 | show "\<exists>s>0. \<forall>z. z\<noteq>0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>g(f x + z) - g(f x)\<bar> < r" | |
| 1444 | proof (intro exI conjI) | |
| 23441 | 1445 | show "0<e'" using e' . | 
| 21164 | 1446 | show "\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < e' \<longrightarrow> \<bar>g (f x + z) - g (f x)\<bar> < r" | 
| 1447 | proof (intro strip) | |
| 1448 | fix z::real | |
| 1449 | assume z: "z \<noteq> 0 \<and> \<bar>z\<bar> < e'" | |
| 1450 | with e e_lt e_simps all [rule_format, of "f x + z"] | |
| 1451 | show "\<bar>g (f x + z) - g (f x)\<bar> < r" by force | |
| 1452 | qed | |
| 1453 | qed | |
| 1454 | qed | |
| 1455 | qed | |
| 1456 | ||
| 23041 | 1457 | text {* Derivative of inverse function *}
 | 
| 1458 | ||
| 1459 | lemma DERIV_inverse_function: | |
| 1460 | fixes f g :: "real \<Rightarrow> real" | |
| 1461 | assumes der: "DERIV f (g x) :> D" | |
| 1462 | assumes neq: "D \<noteq> 0" | |
| 23044 | 1463 | assumes a: "a < x" and b: "x < b" | 
| 1464 | assumes inj: "\<forall>y. a < y \<and> y < b \<longrightarrow> f (g y) = y" | |
| 23041 | 1465 | assumes cont: "isCont g x" | 
| 1466 | shows "DERIV g x :> inverse D" | |
| 1467 | unfolding DERIV_iff2 | |
| 23044 | 1468 | proof (rule LIM_equal2) | 
| 1469 | show "0 < min (x - a) (b - x)" | |
| 27668 | 1470 | using a b by arith | 
| 23044 | 1471 | next | 
| 23041 | 1472 | fix y | 
| 23044 | 1473 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 27668 | 1474 | hence "a < y" and "y < b" | 
| 23044 | 1475 | by (simp_all add: abs_less_iff) | 
| 23041 | 1476 | thus "(g y - g x) / (y - x) = | 
| 1477 | inverse ((f (g y) - x) / (g y - g x))" | |
| 1478 | by (simp add: inj) | |
| 1479 | next | |
| 1480 | have "(\<lambda>z. (f z - f (g x)) / (z - g x)) -- g x --> D" | |
| 1481 | by (rule der [unfolded DERIV_iff2]) | |
| 1482 | hence 1: "(\<lambda>z. (f z - x) / (z - g x)) -- g x --> D" | |
| 23044 | 1483 | using inj a b by simp | 
| 23041 | 1484 | have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x" | 
| 1485 | proof (safe intro!: exI) | |
| 23044 | 1486 | show "0 < min (x - a) (b - x)" | 
| 1487 | using a b by simp | |
| 23041 | 1488 | next | 
| 1489 | fix y | |
| 23044 | 1490 | assume "norm (y - x) < min (x - a) (b - x)" | 
| 1491 | hence y: "a < y" "y < b" | |
| 1492 | by (simp_all add: abs_less_iff) | |
| 23041 | 1493 | assume "g y = g x" | 
| 1494 | hence "f (g y) = f (g x)" by simp | |
| 23044 | 1495 | hence "y = x" using inj y a b by simp | 
| 23041 | 1496 | also assume "y \<noteq> x" | 
| 1497 | finally show False by simp | |
| 1498 | qed | |
| 1499 | have "(\<lambda>y. (f (g y) - x) / (g y - g x)) -- x --> D" | |
| 1500 | using cont 1 2 by (rule isCont_LIM_compose2) | |
| 1501 | thus "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x))) | |
| 1502 | -- x --> inverse D" | |
| 1503 | using neq by (rule LIM_inverse) | |
| 1504 | qed | |
| 1505 | ||
| 29975 | 1506 | |
| 1507 | subsection {* Generalized Mean Value Theorem *}
 | |
| 1508 | ||
| 21164 | 1509 | theorem GMVT: | 
| 21784 
e76faa6e65fd
changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
 huffman parents: 
21404diff
changeset | 1510 | fixes a b :: real | 
| 21164 | 1511 | assumes alb: "a < b" | 
| 1512 | and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x" | |
| 1513 | and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable x" | |
| 1514 | and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x" | |
| 1515 | and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable x" | |
| 1516 | shows "\<exists>g'c f'c c. 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)" | |
| 1517 | proof - | |
| 1518 | let ?h = "\<lambda>x. (f b - f a)*(g x) - (g b - g a)*(f x)" | |
| 1519 | from prems have "a < b" by simp | |
| 1520 | moreover have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?h x" | |
| 1521 | proof - | |
| 1522 | have "\<forall>x. a <= x \<and> x <= b \<longrightarrow> isCont (\<lambda>x. f b - f a) x" by simp | |
| 1523 | with gc have "\<forall>x. a <= x \<and> x <= b \<longrightarrow> isCont (\<lambda>x. (f b - f a) * g x) x" | |
| 1524 | by (auto intro: isCont_mult) | |
| 1525 | moreover | |
| 1526 | have "\<forall>x. a <= x \<and> x <= b \<longrightarrow> isCont (\<lambda>x. g b - g a) x" by simp | |
| 1527 | with fc have "\<forall>x. a <= x \<and> x <= b \<longrightarrow> isCont (\<lambda>x. (g b - g a) * f x) x" | |
| 1528 | by (auto intro: isCont_mult) | |
| 1529 | ultimately show ?thesis | |
| 1530 | by (fastsimp intro: isCont_diff) | |
| 1531 | qed | |
| 1532 | moreover | |
| 1533 | have "\<forall>x. a < x \<and> x < b \<longrightarrow> ?h differentiable x" | |
| 1534 | proof - | |
| 35216 | 1535 | have "\<forall>x. a < x \<and> x < b \<longrightarrow> (\<lambda>x. f b - f a) differentiable x" by simp | 
| 1536 | with gd have "\<forall>x. a < x \<and> x < b \<longrightarrow> (\<lambda>x. (f b - f a) * g x) differentiable x" by simp | |
| 21164 | 1537 | moreover | 
| 35216 | 1538 | have "\<forall>x. a < x \<and> x < b \<longrightarrow> (\<lambda>x. g b - g a) differentiable x" by simp | 
| 1539 | with fd have "\<forall>x. a < x \<and> x < b \<longrightarrow> (\<lambda>x. (g b - g a) * f x) differentiable x" by simp | |
| 1540 | ultimately show ?thesis by simp | |
| 21164 | 1541 | qed | 
| 1542 | ultimately have "\<exists>l z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" by (rule MVT) | |
| 1543 | then obtain l where ldef: "\<exists>z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 1544 | then obtain c where cdef: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" .. | |
| 1545 | ||
| 1546 | from cdef have cint: "a < c \<and> c < b" by auto | |
| 1547 | with gd have "g differentiable c" by simp | |
| 1548 | hence "\<exists>D. DERIV g c :> D" by (rule differentiableD) | |
| 1549 | then obtain g'c where g'cdef: "DERIV g c :> g'c" .. | |
| 1550 | ||
| 1551 | from cdef have "a < c \<and> c < b" by auto | |
| 1552 | with fd have "f differentiable c" by simp | |
| 1553 | hence "\<exists>D. DERIV f c :> D" by (rule differentiableD) | |
| 1554 | then obtain f'c where f'cdef: "DERIV f c :> f'c" .. | |
| 1555 | ||
| 1556 | from cdef have "DERIV ?h c :> l" by auto | |
| 1557 | moreover | |
| 1558 |   {
 | |
| 23441 | 1559 | have "DERIV (\<lambda>x. (f b - f a) * g x) c :> g'c * (f b - f a)" | 
| 21164 | 1560 | apply (insert DERIV_const [where k="f b - f a"]) | 
| 1561 | apply (drule meta_spec [of _ c]) | |
| 23441 | 1562 | apply (drule DERIV_mult [OF _ g'cdef]) | 
| 1563 | by simp | |
| 1564 | moreover have "DERIV (\<lambda>x. (g b - g a) * f x) c :> f'c * (g b - g a)" | |
| 21164 | 1565 | apply (insert DERIV_const [where k="g b - g a"]) | 
| 1566 | apply (drule meta_spec [of _ c]) | |
| 23441 | 1567 | apply (drule DERIV_mult [OF _ f'cdef]) | 
| 1568 | by simp | |
| 21164 | 1569 | ultimately have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" | 
| 1570 | by (simp add: DERIV_diff) | |
| 1571 | } | |
| 1572 | ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) | |
| 1573 | ||
| 1574 |   {
 | |
| 1575 | from cdef have "?h b - ?h a = (b - a) * l" by auto | |
| 1576 | also with leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp | |
| 1577 | finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp | |
| 1578 | } | |
| 1579 | moreover | |
| 1580 |   {
 | |
| 1581 | have "?h b - ?h a = | |
| 1582 | ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - | |
| 1583 | ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" | |
| 29667 | 1584 | by (simp add: algebra_simps) | 
| 21164 | 1585 | hence "?h b - ?h a = 0" by auto | 
| 1586 | } | |
| 1587 | ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto | |
| 1588 | with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp | |
| 1589 | hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp | |
| 1590 | hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac) | |
| 1591 | ||
| 1592 | with g'cdef f'cdef cint show ?thesis by auto | |
| 1593 | qed | |
| 1594 | ||
| 29470 
1851088a1f87
convert Deriv.thy to use new Polynomial library (incomplete)
 huffman parents: 
29169diff
changeset | 1595 | |
| 29166 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1596 | subsection {* Theorems about Limits *}
 | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1597 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1598 | (* need to rename second isCont_inverse *) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1599 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1600 | lemma isCont_inv_fun: | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1601 | fixes f g :: "real \<Rightarrow> real" | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1602 | shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z; | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1603 | \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |] | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1604 | ==> isCont g (f x)" | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1605 | by (rule isCont_inverse_function) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1606 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1607 | lemma isCont_inv_fun_inv: | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1608 | fixes f g :: "real \<Rightarrow> real" | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1609 | shows "[| 0 < d; | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1610 | \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z; | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1611 | \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |] | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1612 | ==> \<exists>e. 0 < e & | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1613 | (\<forall>y. 0 < \<bar>y - f(x)\<bar> & \<bar>y - f(x)\<bar> < e --> f(g(y)) = y)" | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1614 | apply (drule isCont_inj_range) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1615 | prefer 2 apply (assumption, assumption, auto) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1616 | apply (rule_tac x = e in exI, auto) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1617 | apply (rotate_tac 2) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1618 | apply (drule_tac x = y in spec, auto) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1619 | done | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1620 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1621 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1622 | text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
 | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1623 | lemma LIM_fun_gt_zero: | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1624 | "[| f -- c --> (l::real); 0 < l |] | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1625 | ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> 0 < f x)" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 1626 | apply (auto simp add: LIM_eq) | 
| 29166 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1627 | apply (drule_tac x = "l/2" in spec, safe, force) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1628 | apply (rule_tac x = s in exI) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1629 | apply (auto simp only: abs_less_iff) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1630 | done | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1631 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1632 | lemma LIM_fun_less_zero: | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1633 | "[| f -- c --> (l::real); l < 0 |] | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1634 | ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x < 0)" | 
| 31338 
d41a8ba25b67
generalize constants from Lim.thy to class metric_space
 huffman parents: 
31336diff
changeset | 1635 | apply (auto simp add: LIM_eq) | 
| 29166 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1636 | apply (drule_tac x = "-l/2" in spec, safe, force) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1637 | apply (rule_tac x = s in exI) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1638 | apply (auto simp only: abs_less_iff) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1639 | done | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1640 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1641 | |
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1642 | lemma LIM_fun_not_zero: | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1643 | "[| f -- c --> (l::real); l \<noteq> 0 |] | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1644 | ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x \<noteq> 0)" | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1645 | apply (cut_tac x = l and y = 0 in linorder_less_linear, auto) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1646 | apply (drule LIM_fun_less_zero) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1647 | apply (drule_tac [3] LIM_fun_gt_zero) | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1648 | apply force+ | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1649 | done | 
| 
c23b2d108612
move theorems about limits from Transcendental.thy to Deriv.thy
 huffman parents: 
28952diff
changeset | 1650 | |
| 21164 | 1651 | end |