move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
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
8 header{* Differentiation *}
14 text{*Standard Definitions*}
17 deriv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
18 --{*Differentiation: D is derivative of function f at x*}
19 ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
20 "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)"
22 subsection {* Derivatives *}
24 lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)"
25 by (simp add: deriv_def)
27 lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D"
28 by (simp add: deriv_def)
30 lemma DERIV_const [simp]: "DERIV (\<lambda>x. k) x :> 0"
31 by (simp add: deriv_def tendsto_const)
33 lemma DERIV_ident [simp]: "DERIV (\<lambda>x. x) x :> 1"
34 by (simp add: deriv_def tendsto_const cong: LIM_cong)
37 "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + g x) x :> D + E"
38 by (simp only: deriv_def add_diff_add add_divide_distrib tendsto_add)
41 "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. - f x) x :> - D"
42 by (simp only: deriv_def minus_diff_minus divide_minus_left tendsto_minus)
45 "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x - g x) x :> D - E"
46 by (simp only: diff_minus DERIV_add DERIV_minus)
48 lemma DERIV_add_minus:
49 "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + - g x) x :> D + - E"
50 by (simp only: DERIV_add DERIV_minus)
52 lemma DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x"
53 proof (unfold isCont_iff)
54 assume "DERIV f x :> D"
55 hence "(\<lambda>h. (f(x+h) - f(x)) / h) -- 0 --> D"
57 hence "(\<lambda>h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0"
58 by (intro tendsto_mult tendsto_ident_at)
59 hence "(\<lambda>h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0"
61 hence "(\<lambda>h. f(x+h) - f(x)) -- 0 --> 0"
62 by (simp cong: LIM_cong)
63 thus "(\<lambda>h. f(x+h)) -- 0 --> f(x)"
64 by (simp add: LIM_def dist_norm)
67 lemma DERIV_mult_lemma:
68 fixes a b c d :: "'a::real_field"
69 shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d"
70 by (simp add: field_simps diff_divide_distrib)
73 assumes f: "DERIV f x :> D"
74 assumes g: "DERIV g x :> E"
75 shows "DERIV (\<lambda>x. f x * g x) x :> f x * E + D * g x"
76 proof (unfold deriv_def)
77 from f have "isCont f x"
78 by (rule DERIV_isCont)
79 hence "(\<lambda>h. f(x+h)) -- 0 --> f x"
80 by (simp only: isCont_iff)
81 hence "(\<lambda>h. f(x+h) * ((g(x+h) - g x) / h) +
82 ((f(x+h) - f x) / h) * g x)
83 -- 0 --> f x * E + D * g x"
84 by (intro tendsto_intros DERIV_D f g)
85 thus "(\<lambda>h. (f(x+h) * g(x+h) - f x * g x) / h)
86 -- 0 --> f x * E + D * g x"
87 by (simp only: DERIV_mult_lemma)
91 "DERIV f x :> Da \<Longrightarrow> DERIV g x :> Db \<Longrightarrow> DERIV (\<lambda>x. f x * g x) x :> Da * g x + Db * f x"
92 by (drule (1) DERIV_mult', simp only: mult_commute add_commute)
95 "DERIV f x :> D \<Longrightarrow> DERIV f x :> E \<Longrightarrow> D = E"
96 unfolding deriv_def by (rule LIM_unique)
98 text{*Differentiation of finite sum*}
102 and "\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x :> (f' x n)"
103 shows "DERIV (%x. setsum (f x) S) x :> setsum (f' x) S"
104 using assms by induct (auto intro!: DERIV_add)
106 lemma DERIV_sumr [rule_format (no_asm)]:
107 "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x))
108 --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x :> (\<Sum>r=m..<n. f' r x)"
109 by (auto intro: DERIV_setsum)
111 text{*Alternative definition for differentiability*}
114 fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
115 "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) =
116 ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
118 apply (drule_tac k="- a" in LIM_offset)
119 apply (simp add: diff_minus)
120 apply (drule_tac k="a" in LIM_offset)
121 apply (simp add: add_commute)
124 lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)"
125 by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
127 lemma DERIV_inverse_lemma:
128 "\<lbrakk>a \<noteq> 0; b \<noteq> (0::'a::real_normed_field)\<rbrakk>
129 \<Longrightarrow> (inverse a - inverse b) / h
130 = - (inverse a * ((a - b) / h) * inverse b)"
131 by (simp add: inverse_diff_inverse)
133 lemma DERIV_inverse':
134 assumes der: "DERIV f x :> D"
135 assumes neq: "f x \<noteq> 0"
136 shows "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))"
137 (is "DERIV _ _ :> ?E")
138 proof (unfold DERIV_iff2)
139 from der have lim_f: "f -- x --> f x"
140 by (rule DERIV_isCont [unfolded isCont_def])
142 from neq have "0 < norm (f x)" by simp
143 with LIM_D [OF lim_f] obtain s
145 and less_fx: "\<And>z. \<lbrakk>z \<noteq> x; norm (z - x) < s\<rbrakk>
146 \<Longrightarrow> norm (f z - f x) < norm (f x)"
149 show "(\<lambda>z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E"
150 proof (rule LIM_equal2 [OF s])
152 assume "z \<noteq> x" "norm (z - x) < s"
153 hence "norm (f z - f x) < norm (f x)" by (rule less_fx)
154 hence "f z \<noteq> 0" by auto
155 thus "(inverse (f z) - inverse (f x)) / (z - x) =
156 - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))"
157 using neq by (rule DERIV_inverse_lemma)
159 from der have "(\<lambda>z. (f z - f x) / (z - x)) -- x --> D"
160 by (unfold DERIV_iff2)
161 thus "(\<lambda>z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x)))
163 by (intro tendsto_intros lim_f neq)
168 "\<lbrakk>DERIV f x :> D; DERIV g x :> E; g x \<noteq> 0\<rbrakk>
169 \<Longrightarrow> DERIV (\<lambda>x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)"
170 apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) +
171 D * inverse (g x) = (D * g x - f x * E) / (g x * g x)")
173 apply (unfold divide_inverse)
174 apply (erule DERIV_mult')
175 apply (erule (1) DERIV_inverse')
176 apply (simp add: ring_distribs nonzero_inverse_mult_distrib)
179 lemma DERIV_power_Suc:
180 fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
181 assumes f: "DERIV f x :> D"
182 shows "DERIV (\<lambda>x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)"
185 show ?case by (simp add: f)
187 from DERIV_mult' [OF f Suc] show ?case
188 apply (simp only: of_nat_Suc ring_distribs mult_1_left)
189 apply (simp only: power_Suc algebra_simps)
194 fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
195 assumes f: "DERIV f x :> D"
196 shows "DERIV (\<lambda>x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))"
197 by (cases "n", simp, simp add: DERIV_power_Suc f del: power_Suc)
199 text {* Caratheodory formulation of derivative at a point *}
203 (\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isCont g x & g x = l)"
206 assume der: "DERIV f x :> l"
207 show "\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) \<and> isCont g x \<and> g x = l"
208 proof (intro exI conjI)
209 let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
210 show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
211 show "isCont ?g x" using der
212 by (simp add: isCont_iff DERIV_iff diff_minus
213 cong: LIM_equal [rule_format])
214 show "?g x = l" by simp
219 "(\<forall>z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast
220 thus "(DERIV f x :> l)"
221 by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong)
225 assumes f: "DERIV f x :> D"
226 assumes g: "DERIV g (f x) :> E"
227 shows "DERIV (\<lambda>x. g (f x)) x :> E * D"
228 proof (unfold DERIV_iff2)
229 obtain d where d: "\<forall>y. g y - g (f x) = d y * (y - f x)"
230 and cont_d: "isCont d (f x)" and dfx: "d (f x) = E"
231 using CARAT_DERIV [THEN iffD1, OF g] by fast
232 from f have "f -- x --> f x"
233 by (rule DERIV_isCont [unfolded isCont_def])
234 with cont_d have "(\<lambda>z. d (f z)) -- x --> d (f x)"
235 by (rule isCont_tendsto_compose)
236 hence "(\<lambda>z. d (f z) * ((f z - f x) / (z - x)))
237 -- x --> d (f x) * D"
238 by (rule tendsto_mult [OF _ f [unfolded DERIV_iff2]])
239 thus "(\<lambda>z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D"
244 Let's do the standard proof, though theorem
245 @{text "LIM_mult2"} follows from a NS proof
249 "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D"
250 by (drule DERIV_mult' [OF DERIV_const], simp)
252 lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c"
253 apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force)
254 apply (erule DERIV_cmult)
257 text {* Standard version *}
258 lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db"
259 by (drule (1) DERIV_chain', simp add: o_def mult_commute)
261 lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db"
262 by (auto dest: DERIV_chain simp add: o_def)
264 text {* Derivative of linear multiplication *}
265 lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c"
266 by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp)
268 lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
269 apply (cut_tac DERIV_power [OF DERIV_ident])
270 apply (simp add: real_of_nat_def)
273 text {* Power of @{text "-1"} *}
276 fixes x :: "'a::{real_normed_field}"
277 shows "x \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))"
278 by (drule DERIV_inverse' [OF DERIV_ident]) simp
280 text {* Derivative of inverse *}
281 lemma DERIV_inverse_fun:
282 fixes x :: "'a::{real_normed_field}"
283 shows "[| DERIV f x :> d; f(x) \<noteq> 0 |]
284 ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"
285 by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib)
287 text {* Derivative of quotient *}
288 lemma DERIV_quotient:
289 fixes x :: "'a::{real_normed_field}"
290 shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> 0 |]
291 ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
292 by (drule (2) DERIV_divide) (simp add: mult_commute)
294 text {* @{text "DERIV_intros"} *}
296 structure Deriv_Intros = Named_Thms
298 val name = @{binding DERIV_intros}
299 val description = "DERIV introduction rules"
303 setup Deriv_Intros.setup
305 lemma DERIV_cong: "\<lbrakk> DERIV f x :> X ; X = Y \<rbrakk> \<Longrightarrow> DERIV f x :> Y"
309 DERIV_const[THEN DERIV_cong, DERIV_intros]
310 DERIV_ident[THEN DERIV_cong, DERIV_intros]
311 DERIV_add[THEN DERIV_cong, DERIV_intros]
312 DERIV_minus[THEN DERIV_cong, DERIV_intros]
313 DERIV_mult[THEN DERIV_cong, DERIV_intros]
314 DERIV_diff[THEN DERIV_cong, DERIV_intros]
315 DERIV_inverse'[THEN DERIV_cong, DERIV_intros]
316 DERIV_divide[THEN DERIV_cong, DERIV_intros]
317 DERIV_power[where 'a=real, THEN DERIV_cong,
318 unfolded real_of_nat_def[symmetric], DERIV_intros]
319 DERIV_setsum[THEN DERIV_cong, DERIV_intros]
322 subsection {* Differentiability predicate *}
325 differentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
326 (infixl "differentiable" 60) where
327 "f differentiable x = (\<exists>D. DERIV f x :> D)"
329 lemma differentiableE [elim?]:
330 assumes "f differentiable x"
331 obtains df where "DERIV f x :> df"
332 using assms unfolding differentiable_def ..
334 lemma differentiableD: "f differentiable x ==> \<exists>D. DERIV f x :> D"
335 by (simp add: differentiable_def)
337 lemma differentiableI: "DERIV f x :> D ==> f differentiable x"
338 by (force simp add: differentiable_def)
340 lemma differentiable_ident [simp]: "(\<lambda>x. x) differentiable x"
341 by (rule DERIV_ident [THEN differentiableI])
343 lemma differentiable_const [simp]: "(\<lambda>z. a) differentiable x"
344 by (rule DERIV_const [THEN differentiableI])
346 lemma differentiable_compose:
347 assumes f: "f differentiable (g x)"
348 assumes g: "g differentiable x"
349 shows "(\<lambda>x. f (g x)) differentiable x"
351 from `f differentiable (g x)` obtain df where "DERIV f (g x) :> df" ..
353 from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
355 have "DERIV (\<lambda>x. f (g x)) x :> df * dg" by (rule DERIV_chain2)
356 thus ?thesis by (rule differentiableI)
359 lemma differentiable_sum [simp]:
360 assumes "f differentiable x"
361 and "g differentiable x"
362 shows "(\<lambda>x. f x + g x) differentiable x"
364 from `f differentiable x` obtain df where "DERIV f x :> df" ..
366 from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
368 have "DERIV (\<lambda>x. f x + g x) x :> df + dg" by (rule DERIV_add)
369 thus ?thesis by (rule differentiableI)
372 lemma differentiable_minus [simp]:
373 assumes "f differentiable x"
374 shows "(\<lambda>x. - f x) differentiable x"
376 from `f differentiable x` obtain df where "DERIV f x :> df" ..
377 hence "DERIV (\<lambda>x. - f x) x :> - df" by (rule DERIV_minus)
378 thus ?thesis by (rule differentiableI)
381 lemma differentiable_diff [simp]:
382 assumes "f differentiable x"
383 assumes "g differentiable x"
384 shows "(\<lambda>x. f x - g x) differentiable x"
385 unfolding diff_minus using assms by simp
387 lemma differentiable_mult [simp]:
388 assumes "f differentiable x"
389 assumes "g differentiable x"
390 shows "(\<lambda>x. f x * g x) differentiable x"
392 from `f differentiable x` obtain df where "DERIV f x :> df" ..
394 from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
396 have "DERIV (\<lambda>x. f x * g x) x :> df * g x + dg * f x" by (rule DERIV_mult)
397 thus ?thesis by (rule differentiableI)
400 lemma differentiable_inverse [simp]:
401 assumes "f differentiable x" and "f x \<noteq> 0"
402 shows "(\<lambda>x. inverse (f x)) differentiable x"
404 from `f differentiable x` obtain df where "DERIV f x :> df" ..
405 hence "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * df * inverse (f x))"
406 using `f x \<noteq> 0` by (rule DERIV_inverse')
407 thus ?thesis by (rule differentiableI)
410 lemma differentiable_divide [simp]:
411 assumes "f differentiable x"
412 assumes "g differentiable x" and "g x \<noteq> 0"
413 shows "(\<lambda>x. f x / g x) differentiable x"
414 unfolding divide_inverse using assms by simp
416 lemma differentiable_power [simp]:
417 fixes f :: "'a::{real_normed_field} \<Rightarrow> 'a"
418 assumes "f differentiable x"
419 shows "(\<lambda>x. f x ^ n) differentiable x"
422 apply (simp add: assms)
425 subsection {* Local extrema *}
427 text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
429 lemma DERIV_pos_inc_right:
430 fixes f :: "real => real"
431 assumes der: "DERIV f x :> l"
433 shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x + h)"
435 from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
436 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)"
437 by (simp add: diff_minus)
440 and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
443 proof (intro exI conjI strip)
446 assume "0 < h" "h < s"
447 with all [of h] show "f x < f (x+h)"
448 proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
449 split add: split_if_asm)
450 assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
452 have "0 < (f (x+h) - f x) / h" by arith
454 by (simp add: pos_less_divide_eq h)
459 lemma DERIV_neg_dec_left:
460 fixes f :: "real => real"
461 assumes der: "DERIV f x :> l"
463 shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x-h)"
465 from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
466 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)"
467 by (simp add: diff_minus)
470 and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
473 proof (intro exI conjI strip)
476 assume "0 < h" "h < s"
477 with all [of "-h"] show "f x < f (x-h)"
478 proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
479 split add: split_if_asm)
480 assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
482 have "0 < (f (x-h) - f x) / h" by arith
484 by (simp add: pos_less_divide_eq h)
489 lemma DERIV_pos_inc_left:
490 fixes f :: "real => real"
491 shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)"
492 apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified])
493 apply (auto simp add: DERIV_minus)
496 lemma DERIV_neg_dec_right:
497 fixes f :: "real => real"
498 shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x) > f(x + h)"
499 apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified])
500 apply (auto simp add: DERIV_minus)
503 lemma DERIV_local_max:
504 fixes f :: "real => real"
505 assumes der: "DERIV f x :> l"
507 and le: "\<forall>y. \<bar>x-y\<bar> < d --> f(y) \<le> f(x)"
509 proof (cases rule: linorder_cases [of l 0])
510 case equal thus ?thesis .
513 from DERIV_neg_dec_left [OF der less]
514 obtain d' where d': "0 < d'"
515 and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x-h)" by blast
516 from real_lbound_gt_zero [OF d d']
517 obtain e where "0 < e \<and> e < d \<and> e < d'" ..
518 with lt le [THEN spec [where x="x-e"]]
519 show ?thesis by (auto simp add: abs_if)
522 from DERIV_pos_inc_right [OF der greater]
523 obtain d' where d': "0 < d'"
524 and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" by blast
525 from real_lbound_gt_zero [OF d d']
526 obtain e where "0 < e \<and> e < d \<and> e < d'" ..
527 with lt le [THEN spec [where x="x+e"]]
528 show ?thesis by (auto simp add: abs_if)
532 text{*Similar theorem for a local minimum*}
533 lemma DERIV_local_min:
534 fixes f :: "real => real"
535 shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0"
536 by (drule DERIV_minus [THEN DERIV_local_max], auto)
539 text{*In particular, if a function is locally flat*}
540 lemma DERIV_local_const:
541 fixes f :: "real => real"
542 shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0"
543 by (auto dest!: DERIV_local_max)
546 subsection {* Rolle's Theorem *}
548 text{*Lemma about introducing open ball in open interval*}
549 lemma lemma_interval_lt:
551 ==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)"
553 apply (simp add: abs_less_iff)
554 apply (insert linorder_linear [of "x-a" "b-x"], safe)
555 apply (rule_tac x = "x-a" in exI)
556 apply (rule_tac [2] x = "b-x" in exI, auto)
559 lemma lemma_interval: "[| a < x; x < b |] ==>
560 \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a \<le> y & y \<le> b)"
561 apply (drule lemma_interval_lt, auto)
565 text{*Rolle's Theorem.
566 If @{term f} is defined and continuous on the closed interval
567 @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
568 and @{term "f(a) = f(b)"},
569 then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
572 and eq: "f(a) = f(b)"
573 and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
574 and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
575 shows "\<exists>z::real. a < z & z < b & DERIV f z :> 0"
577 have le: "a \<le> b" using lt by simp
578 from isCont_eq_Ub [OF le con]
579 obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x"
580 and alex: "a \<le> x" and xleb: "x \<le> b"
582 from isCont_eq_Lb [OF le con]
583 obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z"
584 and alex': "a \<le> x'" and x'leb: "x' \<le> b"
588 assume axb: "a < x & x < b"
589 --{*@{term f} attains its maximum within the interval*}
590 hence ax: "a<x" and xb: "x<b" by arith +
591 from lemma_interval [OF ax xb]
592 obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
594 hence bound': "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> f y \<le> f x" using x_max
596 from differentiableD [OF dif [OF axb]]
597 obtain l where der: "DERIV f x :> l" ..
598 have "l=0" by (rule DERIV_local_max [OF der d bound'])
599 --{*the derivative at a local maximum is zero*}
600 thus ?thesis using ax xb der by auto
602 assume notaxb: "~ (a < x & x < b)"
603 hence xeqab: "x=a | x=b" using alex xleb by arith
604 hence fb_eq_fx: "f b = f x" by (auto simp add: eq)
607 assume ax'b: "a < x' & x' < b"
608 --{*@{term f} attains its minimum within the interval*}
609 hence ax': "a<x'" and x'b: "x'<b" by arith+
610 from lemma_interval [OF ax' x'b]
611 obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
613 hence bound': "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> f x' \<le> f y" using x'_min
615 from differentiableD [OF dif [OF ax'b]]
616 obtain l where der: "DERIV f x' :> l" ..
617 have "l=0" by (rule DERIV_local_min [OF der d bound'])
618 --{*the derivative at a local minimum is zero*}
619 thus ?thesis using ax' x'b der by auto
621 assume notax'b: "~ (a < x' & x' < b)"
622 --{*@{term f} is constant througout the interval*}
623 hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith
624 hence fb_eq_fx': "f b = f x'" by (auto simp add: eq)
626 obtain r where ar: "a < r" and rb: "r < b" by blast
627 from lemma_interval [OF ar rb]
628 obtain d where d: "0<d" and bound: "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
630 have eq_fb: "\<forall>z. a \<le> z --> z \<le> b --> f z = f b"
633 assume az: "a \<le> z" and zb: "z \<le> b"
635 proof (rule order_antisym)
636 show "f z \<le> f b" by (simp add: fb_eq_fx x_max az zb)
637 show "f b \<le> f z" by (simp add: fb_eq_fx' x'_min az zb)
640 have bound': "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> f r = f y"
643 assume lt: "\<bar>r-y\<bar> < d"
644 hence "f y = f b" by (simp add: eq_fb bound)
645 thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le)
647 from differentiableD [OF dif [OF conjI [OF ar rb]]]
648 obtain l where der: "DERIV f r :> l" ..
649 have "l=0" by (rule DERIV_local_const [OF der d bound'])
650 --{*the derivative of a constant function is zero*}
651 thus ?thesis using ar rb der by auto
657 subsection{*Mean Value Theorem*}
660 "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)"
661 by (cases "a = b") (simp_all add: field_simps)
665 and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
666 and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
667 shows "\<exists>l z::real. a < z & z < b & DERIV f z :> l &
668 (f(b) - f(a) = (b-a) * l)"
670 let ?F = "%x. f x - ((f b - f a) / (b-a)) * x"
671 have contF: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?F x"
672 using con by (fast intro: isCont_intros)
673 have difF: "\<forall>x. a < x \<and> x < b \<longrightarrow> ?F differentiable x"
676 assume ax: "a < x" and xb: "x < b"
677 from differentiableD [OF dif [OF conjI [OF ax xb]]]
678 obtain l where der: "DERIV f x :> l" ..
679 show "?F differentiable x"
680 by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"],
681 blast intro: DERIV_diff DERIV_cmult_Id der)
683 from Rolle [where f = ?F, OF lt lemma_MVT contF difF]
684 obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0"
686 have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)"
687 by (rule DERIV_cmult_Id)
688 hence derF: "DERIV (\<lambda>x. ?F x + (f b - f a) / (b - a) * x) z
689 :> 0 + (f b - f a) / (b - a)"
690 by (rule DERIV_add [OF der])
692 proof (intro exI conjI)
693 show "a < z" using az .
694 show "z < b" using zb .
695 show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp)
696 show "DERIV f z :> ((f b - f a)/(b-a))" using derF by simp
701 "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |]
702 ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))"
704 apply (blast intro: DERIV_isCont)
705 apply (force dest: order_less_imp_le simp add: differentiable_def)
706 apply (blast dest: DERIV_unique order_less_imp_le)
710 text{*A function is constant if its derivative is 0 over an interval.*}
712 lemma DERIV_isconst_end:
713 fixes f :: "real => real"
715 \<forall>x. a \<le> x & x \<le> b --> isCont f x;
716 \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
718 apply (drule MVT, assumption)
719 apply (blast intro: differentiableI)
720 apply (auto dest!: DERIV_unique simp add: diff_eq_eq)
723 lemma DERIV_isconst1:
724 fixes f :: "real => real"
726 \<forall>x. a \<le> x & x \<le> b --> isCont f x;
727 \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
728 ==> \<forall>x. a \<le> x & x \<le> b --> f x = f a"
730 apply (drule_tac x = a in order_le_imp_less_or_eq, safe)
731 apply (drule_tac b = x in DERIV_isconst_end, auto)
734 lemma DERIV_isconst2:
735 fixes f :: "real => real"
737 \<forall>x. a \<le> x & x \<le> b --> isCont f x;
738 \<forall>x. a < x & x < b --> DERIV f x :> 0;
739 a \<le> x; x \<le> b |]
741 apply (blast dest: DERIV_isconst1)
744 lemma DERIV_isconst3: fixes a b x y :: real
745 assumes "a < b" and "x \<in> {a <..< b}" and "y \<in> {a <..< b}"
746 assumes derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
748 proof (cases "x = y")
753 have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0"
754 proof (rule allI, rule impI)
755 fix z :: real assume "?a \<le> z \<and> z \<le> ?b"
756 hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
757 hence "z \<in> {a<..<b}" by auto
758 thus "DERIV f z :> 0" by (rule derivable)
760 hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z"
761 and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto
763 have "?a < ?b" using `x \<noteq> y` by auto
764 from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y]
768 lemma DERIV_isconst_all:
769 fixes f :: "real => real"
770 shows "\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)"
771 apply (rule linorder_cases [of x y])
772 apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+
775 lemma DERIV_const_ratio_const:
776 fixes f :: "real => real"
777 shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k"
778 apply (rule linorder_cases [of a b], auto)
779 apply (drule_tac [!] f = f in MVT)
780 apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
781 apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
784 lemma DERIV_const_ratio_const2:
785 fixes f :: "real => real"
786 shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k"
787 apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1])
788 apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc)
791 lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)"
794 lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)"
797 text{*Gallileo's "trick": average velocity = av. of end velocities*}
799 lemma DERIV_const_average:
800 fixes v :: "real => real"
801 assumes neq: "a \<noteq> (b::real)"
802 and der: "\<forall>x. DERIV v x :> k"
803 shows "v ((a + b)/2) = (v a + v b)/2"
804 proof (cases rule: linorder_cases [of a b])
805 case equal with neq show ?thesis by simp
808 have "(v b - v a) / (b - a) = k"
809 by (rule DERIV_const_ratio_const2 [OF neq der])
810 hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
811 moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k"
812 by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
813 ultimately show ?thesis using neq by force
816 have "(v b - v a) / (b - a) = k"
817 by (rule DERIV_const_ratio_const2 [OF neq der])
818 hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
819 moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k"
820 by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
821 ultimately show ?thesis using neq by (force simp add: add_commute)
824 (* A function with positive derivative is increasing.
825 A simple proof using the MVT, by Jeremy Avigad. And variants.
827 lemma DERIV_pos_imp_increasing:
828 fixes a::real and b::real and f::"real => real"
829 assumes "a < b" and "\<forall>x. a \<le> x & x \<le> b --> (EX y. DERIV f x :> y & y > 0)"
832 assume f: "~ f a < f b"
833 have "EX l z. a < z & z < b & DERIV f z :> l
834 & f b - f a = (b - a) * l"
838 apply (metis DERIV_isCont)
839 apply (metis differentiableI less_le)
841 then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
842 and "f b - f a = (b - a) * l"
844 with assms f have "~(l > 0)"
845 by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le)
846 with assms z show False
847 by (metis DERIV_unique less_le)
850 lemma DERIV_nonneg_imp_nondecreasing:
851 fixes a::real and b::real and f::"real => real"
852 assumes "a \<le> b" and
853 "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<ge> 0)"
854 shows "f a \<le> f b"
855 proof (rule ccontr, cases "a = b")
856 assume "~ f a \<le> f b" and "a = b"
857 then show False by auto
859 assume A: "~ f a \<le> f b"
861 with assms have "EX l z. a < z & z < b & DERIV f z :> l
862 & f b - f a = (b - a) * l"
866 apply (metis DERIV_isCont)
867 apply (metis differentiableI less_le)
869 then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
870 and C: "f b - f a = (b - a) * l"
872 with A have "a < b" "f b < f a" by auto
873 with C have "\<not> l \<ge> 0" by (auto simp add: not_le algebra_simps)
874 (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl)
875 with assms z show False
876 by (metis DERIV_unique order_less_imp_le)
879 lemma DERIV_neg_imp_decreasing:
880 fixes a::real and b::real and f::"real => real"
882 "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y < 0)"
885 have "(%x. -f x) a < (%x. -f x) b"
886 apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"])
889 apply (metis DERIV_minus neg_0_less_iff_less)
895 lemma DERIV_nonpos_imp_nonincreasing:
896 fixes a::real and b::real and f::"real => real"
897 assumes "a \<le> b" and
898 "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<le> 0)"
899 shows "f a \<ge> f b"
901 have "(%x. -f x) a \<le> (%x. -f x) b"
902 apply (rule DERIV_nonneg_imp_nondecreasing [of a b "%x. -f x"])
905 apply (metis DERIV_minus neg_0_le_iff_le)
911 text {* Derivative of inverse function *}
913 lemma DERIV_inverse_function:
914 fixes f g :: "real \<Rightarrow> real"
915 assumes der: "DERIV f (g x) :> D"
916 assumes neq: "D \<noteq> 0"
917 assumes a: "a < x" and b: "x < b"
918 assumes inj: "\<forall>y. a < y \<and> y < b \<longrightarrow> f (g y) = y"
919 assumes cont: "isCont g x"
920 shows "DERIV g x :> inverse D"
922 proof (rule LIM_equal2)
923 show "0 < min (x - a) (b - x)"
927 assume "norm (y - x) < min (x - a) (b - x)"
928 hence "a < y" and "y < b"
929 by (simp_all add: abs_less_iff)
930 thus "(g y - g x) / (y - x) =
931 inverse ((f (g y) - x) / (g y - g x))"
934 have "(\<lambda>z. (f z - f (g x)) / (z - g x)) -- g x --> D"
935 by (rule der [unfolded DERIV_iff2])
936 hence 1: "(\<lambda>z. (f z - x) / (z - g x)) -- g x --> D"
937 using inj a b by simp
938 have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x"
939 proof (safe intro!: exI)
940 show "0 < min (x - a) (b - x)"
944 assume "norm (y - x) < min (x - a) (b - x)"
945 hence y: "a < y" "y < b"
946 by (simp_all add: abs_less_iff)
948 hence "f (g y) = f (g x)" by simp
949 hence "y = x" using inj y a b by simp
950 also assume "y \<noteq> x"
951 finally show False by simp
953 have "(\<lambda>y. (f (g y) - x) / (g y - g x)) -- x --> D"
954 using cont 1 2 by (rule isCont_LIM_compose2)
955 thus "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x)))
957 using neq by (rule tendsto_inverse)
960 subsection {* Generalized Mean Value Theorem *}
965 and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
966 and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable x"
967 and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x"
968 and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable x"
969 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)"
971 let ?h = "\<lambda>x. (f b - f a)*(g x) - (g b - g a)*(f x)"
972 from assms have "a < b" by simp
973 moreover have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?h x"
975 moreover have "\<forall>x. a < x \<and> x < b \<longrightarrow> ?h differentiable x"
977 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)
978 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" ..
979 then obtain c where cdef: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" ..
981 from cdef have cint: "a < c \<and> c < b" by auto
982 with gd have "g differentiable c" by simp
983 hence "\<exists>D. DERIV g c :> D" by (rule differentiableD)
984 then obtain g'c where g'cdef: "DERIV g c :> g'c" ..
986 from cdef have "a < c \<and> c < b" by auto
987 with fd have "f differentiable c" by simp
988 hence "\<exists>D. DERIV f c :> D" by (rule differentiableD)
989 then obtain f'c where f'cdef: "DERIV f c :> f'c" ..
991 from cdef have "DERIV ?h c :> l" by auto
992 moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)"
993 using g'cdef f'cdef by (auto intro!: DERIV_intros)
994 ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique)
997 from cdef have "?h b - ?h a = (b - a) * l" by auto
998 also with leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
999 finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
1004 ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) -
1005 ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))"
1006 by (simp add: algebra_simps)
1007 hence "?h b - ?h a = 0" by auto
1009 ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto
1010 with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp
1011 hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp
1012 hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac)
1014 with g'cdef f'cdef cint show ?thesis by auto
1018 fixes f g :: "real \<Rightarrow> real"
1020 assumes isCont_f: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont f z"
1021 assumes isCont_g: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont g z"
1022 assumes DERIV_g: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV g z :> (g' z)"
1023 assumes DERIV_f: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV f z :> (f' z)"
1024 shows "\<exists>c. a < c \<and> c < b \<and> (f b - f a) * g' c = (g b - g a) * f' c"
1026 have "\<exists>g'c f'c c. DERIV g c :> g'c \<and> DERIV f c :> f'c \<and>
1027 a < c \<and> c < b \<and> (f b - f a) * g'c = (g b - g a) * f'c"
1028 using assms by (intro GMVT) (force simp: differentiable_def)+
1029 then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c"
1030 using DERIV_f DERIV_g by (force dest: DERIV_unique)
1036 subsection {* L'Hopitals rule *}
1038 lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
1039 DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
1040 unfolding DERIV_iff2
1041 proof (rule filterlim_cong)
1042 assume "eventually (\<lambda>x. f x = g x) (nhds x)"
1043 moreover then have "f x = g x" by (auto simp: eventually_nhds)
1044 moreover assume "x = y" "u = v"
1045 ultimately show "eventually (\<lambda>xa. (f xa - f x) / (xa - x) = (g xa - g y) / (xa - y)) (at x)"
1046 by (auto simp: eventually_within at_def elim: eventually_elim1)
1050 "(DERIV f (x + z) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (x + z)) x :> y)"
1051 by (simp add: DERIV_iff field_simps)
1054 "(DERIV f (- x) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (- x::real) :: real) x :> - y)"
1055 by (simp add: deriv_def filterlim_at_split filterlim_at_left_to_right
1056 tendsto_minus_cancel_left field_simps conj_commute)
1058 lemma lhopital_right_0:
1059 fixes f0 g0 :: "real \<Rightarrow> real"
1060 assumes f_0: "(f0 ---> 0) (at_right 0)"
1061 assumes g_0: "(g0 ---> 0) (at_right 0)"
1063 "eventually (\<lambda>x. g0 x \<noteq> 0) (at_right 0)"
1064 "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)"
1065 "eventually (\<lambda>x. DERIV f0 x :> f' x) (at_right 0)"
1066 "eventually (\<lambda>x. DERIV g0 x :> g' x) (at_right 0)"
1067 assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)"
1068 shows "((\<lambda> x. f0 x / g0 x) ---> x) (at_right 0)"
1070 def f \<equiv> "\<lambda>x. if x \<le> 0 then 0 else f0 x"
1071 then have "f 0 = 0" by simp
1073 def g \<equiv> "\<lambda>x. if x \<le> 0 then 0 else g0 x"
1074 then have "g 0 = 0" by simp
1076 have "eventually (\<lambda>x. g0 x \<noteq> 0 \<and> g' x \<noteq> 0 \<and>
1077 DERIV f0 x :> (f' x) \<and> DERIV g0 x :> (g' x)) (at_right 0)"
1078 using ev by eventually_elim auto
1079 then obtain a where [arith]: "0 < a"
1080 and g0_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g0 x \<noteq> 0"
1081 and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0"
1082 and f0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV f0 x :> (f' x)"
1083 and g0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV g0 x :> (g' x)"
1084 unfolding eventually_within eventually_at by (auto simp: dist_real_def)
1086 have g_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g x \<noteq> 0"
1087 using g0_neq_0 by (simp add: g_def)
1089 { fix x assume x: "0 < x" "x < a" then have "DERIV f x :> (f' x)"
1090 by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]])
1091 (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) }
1094 { fix x assume x: "0 < x" "x < a" then have "DERIV g x :> (g' x)"
1095 by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]])
1096 (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) }
1100 using tendsto_const[of "0::real" "at 0"] f_0
1101 unfolding isCont_def f_def
1102 by (intro filterlim_split_at_real)
1103 (auto elim: eventually_elim1
1104 simp add: filterlim_def le_filter_def eventually_within eventually_filtermap)
1107 using tendsto_const[of "0::real" "at 0"] g_0
1108 unfolding isCont_def g_def
1109 by (intro filterlim_split_at_real)
1110 (auto elim: eventually_elim1
1111 simp add: filterlim_def le_filter_def eventually_within eventually_filtermap)
1113 have "\<exists>\<zeta>. \<forall>x\<in>{0 <..< a}. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)"
1114 proof (rule bchoice, rule)
1115 fix x assume "x \<in> {0 <..< a}"
1116 then have x[arith]: "0 < x" "x < a" by auto
1117 with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x"
1119 have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont f x"
1120 using `isCont f 0` f by (auto intro: DERIV_isCont simp: le_less)
1121 moreover have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont g x"
1122 using `isCont g 0` g by (auto intro: DERIV_isCont simp: le_less)
1123 ultimately have "\<exists>c. 0 < c \<and> c < x \<and> (f x - f 0) * g' c = (g x - g 0) * f' c"
1124 using f g `x < a` by (intro GMVT') auto
1127 with g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c"
1128 by (simp add: field_simps)
1129 ultimately show "\<exists>y. 0 < y \<and> y < x \<and> f x / g x = f' y / g' y"
1130 using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c])
1132 then guess \<zeta> ..
1133 then have \<zeta>: "eventually (\<lambda>x. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)) (at_right 0)"
1134 unfolding eventually_within eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def)
1136 from \<zeta> have "eventually (\<lambda>x. norm (\<zeta> x) \<le> x) (at_right 0)"
1137 by eventually_elim auto
1138 then have "((\<lambda>x. norm (\<zeta> x)) ---> 0) (at_right 0)"
1139 by (rule_tac real_tendsto_sandwich[where f="\<lambda>x. 0" and h="\<lambda>x. x"])
1140 (auto intro: tendsto_const tendsto_ident_at_within)
1141 then have "(\<zeta> ---> 0) (at_right 0)"
1142 by (rule tendsto_norm_zero_cancel)
1143 with \<zeta> have "filterlim \<zeta> (at_right 0) (at_right 0)"
1144 by (auto elim!: eventually_elim1 simp: filterlim_within filterlim_at)
1145 from this lim have "((\<lambda>t. f' (\<zeta> t) / g' (\<zeta> t)) ---> x) (at_right 0)"
1146 by (rule_tac filterlim_compose[of _ _ _ \<zeta>])
1147 ultimately have "((\<lambda>t. f t / g t) ---> x) (at_right 0)" (is ?P)
1148 by (rule_tac filterlim_cong[THEN iffD1, OF refl refl])
1149 (auto elim: eventually_elim1)
1150 also have "?P \<longleftrightarrow> ?thesis"
1151 by (rule filterlim_cong) (auto simp: f_def g_def eventually_within)
1152 finally show ?thesis .
1155 lemma lhopital_right:
1156 "((f::real \<Rightarrow> real) ---> 0) (at_right x) \<Longrightarrow> (g ---> 0) (at_right x) \<Longrightarrow>
1157 eventually (\<lambda>x. g x \<noteq> 0) (at_right x) \<Longrightarrow>
1158 eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow>
1159 eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow>
1160 eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow>
1161 ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow>
1162 ((\<lambda> x. f x / g x) ---> y) (at_right x)"
1163 unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift
1164 by (rule lhopital_right_0)
1166 lemma lhopital_left:
1167 "((f::real \<Rightarrow> real) ---> 0) (at_left x) \<Longrightarrow> (g ---> 0) (at_left x) \<Longrightarrow>
1168 eventually (\<lambda>x. g x \<noteq> 0) (at_left x) \<Longrightarrow>
1169 eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow>
1170 eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow>
1171 eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow>
1172 ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow>
1173 ((\<lambda> x. f x / g x) ---> y) (at_left x)"
1174 unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror
1175 by (rule lhopital_right[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror)
1178 "((f::real \<Rightarrow> real) ---> 0) (at x) \<Longrightarrow> (g ---> 0) (at x) \<Longrightarrow>
1179 eventually (\<lambda>x. g x \<noteq> 0) (at x) \<Longrightarrow>
1180 eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow>
1181 eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow>
1182 eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow>
1183 ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow>
1184 ((\<lambda> x. f x / g x) ---> y) (at x)"
1185 unfolding eventually_at_split filterlim_at_split
1186 by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f'])
1188 lemma lhopital_right_0_at_top:
1189 fixes f g :: "real \<Rightarrow> real"
1190 assumes g_0: "LIM x at_right 0. g x :> at_top"
1192 "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)"
1193 "eventually (\<lambda>x. DERIV f x :> f' x) (at_right 0)"
1194 "eventually (\<lambda>x. DERIV g x :> g' x) (at_right 0)"
1195 assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)"
1196 shows "((\<lambda> x. f x / g x) ---> x) (at_right 0)"
1197 unfolding tendsto_iff
1199 fix e :: real assume "0 < e"
1201 with lim[unfolded tendsto_iff, rule_format, of "e / 4"]
1202 have "eventually (\<lambda>t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp
1203 from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]]
1204 obtain a where [arith]: "0 < a"
1205 and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0"
1206 and f0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV f x :> (f' x)"
1207 and g0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV g x :> (g' x)"
1208 and Df: "\<And>t. 0 < t \<Longrightarrow> t < a \<Longrightarrow> dist (f' t / g' t) x < e / 4"
1209 unfolding eventually_within_le by (auto simp: dist_real_def)
1212 "eventually (\<lambda>t. t < a) (at_right 0)" "eventually (\<lambda>t::real. 0 < t) (at_right 0)"
1213 unfolding eventually_within eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def)
1216 have "eventually (\<lambda>t. 0 < g t) (at_right 0)" "eventually (\<lambda>t. g a < g t) (at_right 0)"
1217 using g_0 by (auto elim: eventually_elim1 simp: filterlim_at_top_dense)
1220 have inv_g: "((\<lambda>x. inverse (g x)) ---> 0) (at_right 0)"
1221 using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl]
1222 by (rule filterlim_compose)
1223 then have "((\<lambda>x. norm (1 - g a * inverse (g x))) ---> norm (1 - g a * 0)) (at_right 0)"
1224 by (intro tendsto_intros)
1225 then have "((\<lambda>x. norm (1 - g a / g x)) ---> 1) (at_right 0)"
1226 by (simp add: inverse_eq_divide)
1227 from this[unfolded tendsto_iff, rule_format, of 1]
1228 have "eventually (\<lambda>x. norm (1 - g a / g x) < 2) (at_right 0)"
1229 by (auto elim!: eventually_elim1 simp: dist_real_def)
1232 from inv_g have "((\<lambda>t. norm ((f a - x * g a) * inverse (g t))) ---> norm ((f a - x * g a) * 0)) (at_right 0)"
1233 by (intro tendsto_intros)
1234 then have "((\<lambda>t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)"
1235 by (simp add: inverse_eq_divide)
1236 from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e`
1237 have "eventually (\<lambda>t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)"
1238 by (auto simp: dist_real_def)
1240 ultimately show "eventually (\<lambda>t. dist (f t / g t) x < e) (at_right 0)"
1241 proof eventually_elim
1242 fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t"
1243 assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2"
1245 have "\<exists>y. t < y \<and> y < a \<and> (g a - g t) * f' y = (f a - f t) * g' y"
1246 using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+
1249 have [arith]: "t < y" "y < a" and D_eq: "(f t - f a) / (g t - g a) = f' y / g' y"
1250 using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps)
1252 have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t"
1253 by (simp add: field_simps)
1254 have "norm (f t / g t - x) \<le>
1255 norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)"
1256 unfolding * by (rule norm_triangle_ineq)
1257 also have "\<dots> = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)"
1258 by (simp add: abs_mult D_eq dist_real_def)
1259 also have "\<dots> < (e / 4) * 2 + e / 2"
1260 using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto
1261 finally show "dist (f t / g t) x < e"
1262 by (simp add: dist_real_def)
1266 lemma lhopital_right_at_top:
1267 "LIM x at_right x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
1268 eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow>
1269 eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow>
1270 eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow>
1271 ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow>
1272 ((\<lambda> x. f x / g x) ---> y) (at_right x)"
1273 unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift
1274 by (rule lhopital_right_0_at_top)
1276 lemma lhopital_left_at_top:
1277 "LIM x at_left x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
1278 eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow>
1279 eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow>
1280 eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow>
1281 ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow>
1282 ((\<lambda> x. f x / g x) ---> y) (at_left x)"
1283 unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror
1284 by (rule lhopital_right_at_top[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror)
1286 lemma lhopital_at_top:
1287 "LIM x at x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
1288 eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow>
1289 eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow>
1290 eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow>
1291 ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow>
1292 ((\<lambda> x. f x / g x) ---> y) (at x)"
1293 unfolding eventually_at_split filterlim_at_split
1294 by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f'])
1296 lemma lhospital_at_top_at_top:
1297 fixes f g :: "real \<Rightarrow> real"
1298 assumes g_0: "LIM x at_top. g x :> at_top"
1299 assumes g': "eventually (\<lambda>x. g' x \<noteq> 0) at_top"
1300 assumes Df: "eventually (\<lambda>x. DERIV f x :> f' x) at_top"
1301 assumes Dg: "eventually (\<lambda>x. DERIV g x :> g' x) at_top"
1302 assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) at_top"
1303 shows "((\<lambda> x. f x / g x) ---> x) at_top"
1304 unfolding filterlim_at_top_to_right
1305 proof (rule lhopital_right_0_at_top)
1306 let ?F = "\<lambda>x. f (inverse x)"
1307 let ?G = "\<lambda>x. g (inverse x)"
1308 let ?R = "at_right (0::real)"
1309 let ?D = "\<lambda>f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))"
1311 show "LIM x ?R. ?G x :> at_top"
1312 using g_0 unfolding filterlim_at_top_to_right .
1314 show "eventually (\<lambda>x. DERIV ?G x :> ?D g' x) ?R"
1315 unfolding eventually_at_right_to_top
1316 using Dg eventually_ge_at_top[where c="1::real"]
1317 apply eventually_elim
1318 apply (rule DERIV_cong)
1319 apply (rule DERIV_chain'[where f=inverse])
1320 apply (auto intro!: DERIV_inverse)
1323 show "eventually (\<lambda>x. DERIV ?F x :> ?D f' x) ?R"
1324 unfolding eventually_at_right_to_top
1325 using Df eventually_ge_at_top[where c="1::real"]
1326 apply eventually_elim
1327 apply (rule DERIV_cong)
1328 apply (rule DERIV_chain'[where f=inverse])
1329 apply (auto intro!: DERIV_inverse)
1332 show "eventually (\<lambda>x. ?D g' x \<noteq> 0) ?R"
1333 unfolding eventually_at_right_to_top
1334 using g' eventually_ge_at_top[where c="1::real"]
1335 by eventually_elim auto
1337 show "((\<lambda>x. ?D f' x / ?D g' x) ---> x) ?R"
1338 unfolding filterlim_at_right_to_top
1339 apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim])
1340 using eventually_ge_at_top[where c="1::real"]
1341 by eventually_elim simp