# HG changeset patch # User hoelzl # Date 1354555147 -3600 # Node ID bbea2e82871c889a598c76503324aed87039203d # Parent b5afeccab2db37873236009d2d3b19158207851a add L'Hôpital's rule diff -r b5afeccab2db -r bbea2e82871c src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Mon Dec 03 18:19:05 2012 +0100 +++ b/src/HOL/Deriv.thy Mon Dec 03 18:19:07 2012 +0100 @@ -1589,4 +1589,183 @@ apply (drule (1) LIM_fun_gt_zero, force) done +lemma GMVT': + fixes f g :: "real \ real" + assumes "a < b" + assumes isCont_f: "\z. a \ z \ z \ b \ isCont f z" + assumes isCont_g: "\z. a \ z \ z \ b \ isCont g z" + assumes DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" + assumes DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" + shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" +proof - + have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ + a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" + using assms by (intro GMVT) (force simp: differentiable_def)+ + then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" + using DERIV_f DERIV_g by (force dest: DERIV_unique) + then show ?thesis + by auto +qed + +lemma lhopital_right_0: + fixes f g :: "real \ real" + assumes f_0: "isCont f 0" "f 0 = 0" + assumes g_0: "isCont g 0" "g 0 = 0" + assumes ev: + "eventually (\x. g x \ 0) (at_right 0)" + "eventually (\x. g' x \ 0) (at_right 0)" + "eventually (\x. DERIV f x :> f' x) (at_right 0)" + "eventually (\x. DERIV g x :> g' x) (at_right 0)" + assumes lim: "((\ x. (f' x / g' x)) ---> x) (at_right 0)" + shows "((\ x. f x / g x) ---> x) (at_right 0)" +proof - + from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) ev(4)]] + obtain a where [arith]: "0 < a" + and g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" + and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" + and f: "\x. 0 < x \ x < a \ DERIV f x :> (f' x)" + and g: "\x. 0 < x \ x < a \ DERIV g x :> (g' x)" + unfolding eventually_within eventually_at by (auto simp: dist_real_def) + + { fix x assume x: "0 \ x" "x < a" + from `0 \ x` have "isCont f x" + unfolding le_less + proof + assume "0 = x" with `isCont f 0` show "isCont f x" by simp + next + assume "0 < x" with f x show ?thesis by (auto intro!: DERIV_isCont) + qed } + note isCont_f = this + + { fix x assume x: "0 \ x" "x < a" + from `0 \ x` have "isCont g x" + unfolding le_less + proof + assume "0 = x" with `isCont g 0` show "isCont g x" by simp + next + assume "0 < x" with g x show ?thesis by (auto intro!: DERIV_isCont) + qed } + note isCont_g = this + + have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" + proof (rule bchoice, rule) + fix x assume "x \ {0 <..< a}" + then have x[arith]: "0 < x" "x < a" by auto + with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" + by auto + + have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" + using isCont_f isCont_g f g `x < a` by (intro GMVT') auto + then guess c .. + moreover + with g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" + by (simp add: field_simps) + ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" + using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c]) + qed + then guess \ .. + then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" + unfolding eventually_within eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) + moreover + from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" + by eventually_elim auto + then have "((\x. norm (\ x)) ---> 0) (at_right 0)" + by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) + (auto intro: tendsto_const tendsto_ident_at_within) + then have "(\ ---> 0) (at_right 0)" + by (rule tendsto_norm_zero_cancel) + with \ have "filterlim \ (at_right 0) (at_right 0)" + by (auto elim!: eventually_elim1 simp: filterlim_within filterlim_at) + from this lim have "((\t. f' (\ t) / g' (\ t)) ---> x) (at_right 0)" + by (rule_tac filterlim_compose[of _ _ _ \]) + ultimately show "((\t. f t / g t) ---> x) (at_right 0)" + apply (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) + apply (erule_tac eventually_elim1) + apply simp_all + done +qed + +lemma lhopital_right_0_at_top: + fixes f g :: "real \ real" + assumes g_0: "LIM x at_right 0. g x :> at_top" + assumes ev: + "eventually (\x. g' x \ 0) (at_right 0)" + "eventually (\x. DERIV f x :> f' x) (at_right 0)" + "eventually (\x. DERIV g x :> g' x) (at_right 0)" + assumes lim: "((\ x. (f' x / g' x)) ---> x) (at_right 0)" + shows "((\ x. f x / g x) ---> x) (at_right 0)" + unfolding tendsto_iff +proof safe + fix e :: real assume "0 < e" + + with lim[unfolded tendsto_iff, rule_format, of "e / 4"] + have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp + from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] + obtain a where [arith]: "0 < a" + and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" + and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" + and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" + and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" + unfolding eventually_within_le by (auto simp: dist_real_def) + + from Df have + "eventually (\t. t < a) (at_right 0)" + "eventually (\t::real. 0 < t) (at_right 0)" + unfolding eventually_within eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) + moreover + + have "eventually (\t. 0 < g t) (at_right 0)" + using g_0[unfolded filterlim_at_top, rule_format, of "1"] by eventually_elim auto + + moreover + + have "eventually (\t. g a < g t) (at_right 0)" + using g_0[unfolded filterlim_at_top, rule_format, of "g a + 1"] by eventually_elim auto + moreover + have inv_g: "((\x. inverse (g x)) ---> 0) (at_right 0)" + using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] + by (rule filterlim_compose) + then have "((\x. norm (1 - g a * inverse (g x))) ---> norm (1 - g a * 0)) (at_right 0)" + by (intro tendsto_intros) + then have "((\x. norm (1 - g a / g x)) ---> 1) (at_right 0)" + by (simp add: inverse_eq_divide) + from this[unfolded tendsto_iff, rule_format, of 1] + have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" + by (auto elim!: eventually_elim1 simp: dist_real_def) + + moreover + from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) ---> norm ((f a - x * g a) * 0)) (at_right 0)" + by (intro tendsto_intros) + then have "((\t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)" + by (simp add: inverse_eq_divide) + from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e` + have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" + by (auto simp: dist_real_def) + + ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" + proof eventually_elim + fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" + assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" + + have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" + using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ + then guess y .. + from this + have [arith]: "t < y" "y < a" and D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" + using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps) + + 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" + by (simp add: field_simps) + have "norm (f t / g t - x) \ + norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" + unfolding * by (rule norm_triangle_ineq) + also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" + by (simp add: abs_mult D_eq dist_real_def) + also have "\ < (e / 4) * 2 + e / 2" + using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto + finally show "dist (f t / g t) x < e" + by (simp add: dist_real_def) + qed +qed + end diff -r b5afeccab2db -r bbea2e82871c src/HOL/Limits.thy --- a/src/HOL/Limits.thy Mon Dec 03 18:19:05 2012 +0100 +++ b/src/HOL/Limits.thy Mon Dec 03 18:19:07 2012 +0100 @@ -445,6 +445,14 @@ shows "eventually P (at a) \ (\d>0. \x. x \ a \ dist x a < d \ P x)" unfolding at_def eventually_within eventually_nhds_metric by auto +lemma eventually_within_less: (* COPY FROM Topo/eventually_within *) + "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a < d \ P x)" + unfolding eventually_within eventually_at dist_nz by auto + +lemma eventually_within_le: (* COPY FROM Topo/eventually_within_le *) + "eventually P (at a within S) \ (\d>0. \x\S. 0 < dist x a \ dist x a <= d \ P x)" + unfolding eventually_within_less by auto (metis dense order_le_less_trans) + lemma at_eq_bot_iff: "at a = bot \ open {a}" unfolding trivial_limit_def eventually_at_topological by (safe, case_tac "S = {a}", simp, fast, fast) @@ -675,14 +683,18 @@ "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" unfolding filterlim_def le_filter_def eventually_filtermap .. -lemma filterlim_compose: +lemma filterlim_compose: "filterlim g F3 F2 \ filterlim f F2 F1 \ filterlim (\x. g (f x)) F3 F1" unfolding filterlim_def filtermap_filtermap[symmetric] by (metis filtermap_mono order_trans) -lemma filterlim_mono: +lemma filterlim_mono: "filterlim f F2 F1 \ F2 \ F2' \ F1' \ F1 \ filterlim f F2' F1'" unfolding filterlim_def by (metis filtermap_mono order_trans) +lemma filterlim_cong: + "F1 = F1' \ F2 = F2' \ eventually (\x. f x = g x) F2 \ filterlim f F1 F2 = filterlim g F1' F2'" + by (auto simp: filterlim_def le_filter_def eventually_filtermap elim: eventually_elim2) + lemma filterlim_within: "(LIM x F1. f x :> F2 within S) \ (eventually (\x. f x \ S) F1 \ (LIM x F1. f x :> F2))" unfolding filterlim_def