diff -r 66c3a7589de7 -r 2d2f59e6055a src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Tue Mar 26 12:20:59 2013 +0100 +++ b/src/HOL/Deriv.thy Tue Mar 26 12:21:00 2013 +0100 @@ -422,178 +422,6 @@ apply (simp add: assms) done - -subsection {* Nested Intervals and Bisection *} - -lemma nested_sequence_unique: - assumes "\n. f n \ f (Suc n)" "\n. g (Suc n) \ g n" "\n. f n \ g n" "(\n. f n - g n) ----> 0" - shows "\l::real. ((\n. f n \ l) \ f ----> l) \ ((\n. l \ g n) \ g ----> l)" -proof - - have "incseq f" unfolding incseq_Suc_iff by fact - have "decseq g" unfolding decseq_Suc_iff by fact - - { fix n - from `decseq g` have "g n \ g 0" by (rule decseqD) simp - with `\n. f n \ g n`[THEN spec, of n] have "f n \ g 0" by auto } - then obtain u where "f ----> u" "\i. f i \ u" - using incseq_convergent[OF `incseq f`] by auto - moreover - { fix n - from `incseq f` have "f 0 \ f n" by (rule incseqD) simp - with `\n. f n \ g n`[THEN spec, of n] have "f 0 \ g n" by simp } - then obtain l where "g ----> l" "\i. l \ g i" - using decseq_convergent[OF `decseq g`] by auto - moreover note LIMSEQ_unique[OF assms(4) tendsto_diff[OF `f ----> u` `g ----> l`]] - ultimately show ?thesis by auto -qed - -lemma Bolzano[consumes 1, case_names trans local]: - fixes P :: "real \ real \ bool" - assumes [arith]: "a \ b" - assumes trans: "\a b c. \P a b; P b c; a \ b; b \ c\ \ P a c" - assumes local: "\x. a \ x \ x \ b \ \d>0. \a b. a \ x \ x \ b \ b - a < d \ P a b" - shows "P a b" -proof - - def bisect \ "nat_rec (a, b) (\n (x, y). if P x ((x+y) / 2) then ((x+y)/2, y) else (x, (x+y)/2))" - def l \ "\n. fst (bisect n)" and u \ "\n. snd (bisect n)" - have l[simp]: "l 0 = a" "\n. l (Suc n) = (if P (l n) ((l n + u n) / 2) then (l n + u n) / 2 else l n)" - and u[simp]: "u 0 = b" "\n. u (Suc n) = (if P (l n) ((l n + u n) / 2) then u n else (l n + u n) / 2)" - by (simp_all add: l_def u_def bisect_def split: prod.split) - - { fix n have "l n \ u n" by (induct n) auto } note this[simp] - - have "\x. ((\n. l n \ x) \ l ----> x) \ ((\n. x \ u n) \ u ----> x)" - proof (safe intro!: nested_sequence_unique) - fix n show "l n \ l (Suc n)" "u (Suc n) \ u n" by (induct n) auto - next - { fix n have "l n - u n = (a - b) / 2^n" by (induct n) (auto simp: field_simps) } - then show "(\n. l n - u n) ----> 0" by (simp add: LIMSEQ_divide_realpow_zero) - qed fact - then obtain x where x: "\n. l n \ x" "\n. x \ u n" and "l ----> x" "u ----> x" by auto - obtain d where "0 < d" and d: "\a b. a \ x \ x \ b \ b - a < d \ P a b" - using `l 0 \ x` `x \ u 0` local[of x] by auto - - show "P a b" - proof (rule ccontr) - assume "\ P a b" - { fix n have "\ P (l n) (u n)" - proof (induct n) - case (Suc n) with trans[of "l n" "(l n + u n) / 2" "u n"] show ?case by auto - qed (simp add: `\ P a b`) } - moreover - { have "eventually (\n. x - d / 2 < l n) sequentially" - using `0 < d` `l ----> x` by (intro order_tendstoD[of _ x]) auto - moreover have "eventually (\n. u n < x + d / 2) sequentially" - using `0 < d` `u ----> x` by (intro order_tendstoD[of _ x]) auto - ultimately have "eventually (\n. P (l n) (u n)) sequentially" - proof eventually_elim - fix n assume "x - d / 2 < l n" "u n < x + d / 2" - from add_strict_mono[OF this] have "u n - l n < d" by simp - with x show "P (l n) (u n)" by (rule d) - qed } - ultimately show False by simp - qed -qed - -(*HOL style here: object-level formulations*) -lemma IVT_objl: "(f(a::real) \ (y::real) & y \ f(b) & a \ b & - (\x. a \ x & x \ b --> isCont f x)) - --> (\x. a \ x & x \ b & f(x) = y)" -apply (blast intro: IVT) -done - -lemma IVT2_objl: "(f(b::real) \ (y::real) & y \ f(a) & a \ b & - (\x. a \ x & x \ b --> isCont f x)) - --> (\x. a \ x & x \ b & f(x) = y)" -apply (blast intro: IVT2) -done - - -lemma compact_Icc[simp, intro]: "compact {a .. b::real}" -proof (cases "a \ b", rule compactI) - fix C assume C: "a \ b" "\t\C. open t" "{a..b} \ \C" - def T == "{a .. b}" - from C(1,3) show "\C'\C. finite C' \ {a..b} \ \C'" - proof (induct rule: Bolzano) - case (trans a b c) - then have *: "{a .. c} = {a .. b} \ {b .. c}" by auto - from trans obtain C1 C2 where "C1\C \ finite C1 \ {a..b} \ \C1" "C2\C \ finite C2 \ {b..c} \ \C2" - by (auto simp: *) - with trans show ?case - unfolding * by (intro exI[of _ "C1 \ C2"]) auto - next - case (local x) - then have "x \ \C" using C by auto - with C(2) obtain c where "x \ c" "open c" "c \ C" by auto - then obtain e where "0 < e" "{x - e <..< x + e} \ c" - by (auto simp: open_real_def dist_real_def subset_eq Ball_def abs_less_iff) - with `c \ C` show ?case - by (safe intro!: exI[of _ "e/2"] exI[of _ "{c}"]) auto - qed -qed simp - -subsection {* Boundedness of continuous functions *} - -text{*By bisection, function continuous on closed interval is bounded above*} - -lemma isCont_eq_Ub: - fixes f :: "real \ 'a::linorder_topology" - shows "a \ b \ \x::real. a \ x \ x \ b \ isCont f x \ - \M. (\x. a \ x \ x \ b \ f x \ M) \ (\x. a \ x \ x \ b \ f x = M)" - using continuous_attains_sup[of "{a .. b}" f] - apply (simp add: continuous_at_imp_continuous_on Ball_def) - apply safe - apply (rule_tac x="f x" in exI) - apply auto - done - -lemma isCont_eq_Lb: - fixes f :: "real \ 'a::linorder_topology" - shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ - \M. (\x. a \ x \ x \ b \ M \ f x) \ (\x. a \ x \ x \ b \ f x = M)" - using continuous_attains_inf[of "{a .. b}" f] - apply (simp add: continuous_at_imp_continuous_on Ball_def) - apply safe - apply (rule_tac x="f x" in exI) - apply auto - done - -lemma isCont_bounded: - fixes f :: "real \ 'a::linorder_topology" - shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ \M. \x. a \ x \ x \ b \ f x \ M" - using isCont_eq_Ub[of a b f] by auto - -lemma isCont_has_Ub: - fixes f :: "real \ 'a::linorder_topology" - shows "a \ b \ \x. a \ x \ x \ b \ isCont f x \ - \M. (\x. a \ x \ x \ b \ f x \ M) \ (\N. N < M \ (\x. a \ x \ x \ b \ N < f x))" - using isCont_eq_Ub[of a b f] by auto - -text{*Refine the above to existence of least upper bound*} - -lemma lemma_reals_complete: "((\x. x \ S) & (\y. isUb UNIV S (y::real))) --> - (\t. isLub UNIV S t)" -by (blast intro: reals_complete) - - -text{*Another version.*} - -lemma isCont_Lb_Ub: "[|a \ b; \x. a \ x & x \ b --> isCont f x |] - ==> \L M::real. (\x::real. a \ x & x \ b --> L \ f(x) & f(x) \ M) & - (\y. L \ y & y \ M --> (\x. a \ x & x \ b & (f(x) = y)))" -apply (frule isCont_eq_Lb) -apply (frule_tac [2] isCont_eq_Ub) -apply (assumption+, safe) -apply (rule_tac x = "f x" in exI) -apply (rule_tac x = "f xa" in exI, simp, safe) -apply (cut_tac x = x and y = xa in linorder_linear, safe) -apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl) -apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe) -apply (rule_tac [2] x = xb in exI) -apply (rule_tac [4] x = xb in exI, simp_all) -done - - subsection {* Local extrema *} text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} @@ -658,7 +486,6 @@ qed qed - lemma DERIV_pos_inc_left: fixes f :: "real => real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d --> f(x - h) < f(x)" @@ -1081,47 +908,6 @@ by simp qed -text{*Continuity of inverse function*} - -lemma isCont_inverse_function: - fixes f g :: "real \ real" - assumes d: "0 < d" - and inj: "\z. \z-x\ \ d \ g (f z) = z" - and cont: "\z. \z-x\ \ d \ isCont f z" - shows "isCont g (f x)" -proof - - let ?A = "f (x - d)" and ?B = "f (x + d)" and ?D = "{x - d..x + d}" - - have f: "continuous_on ?D f" - using cont by (intro continuous_at_imp_continuous_on ballI) auto - then have g: "continuous_on (f`?D) g" - using inj by (intro continuous_on_inv) auto - - from d f have "{min ?A ?B <..< max ?A ?B} \ f ` ?D" - by (intro connected_contains_Ioo connected_continuous_image) (auto split: split_min split_max) - with g have "continuous_on {min ?A ?B <..< max ?A ?B} g" - by (rule continuous_on_subset) - moreover - have "(?A < f x \ f x < ?B) \ (?B < f x \ f x < ?A)" - using d inj by (intro continuous_inj_imp_mono[OF _ _ f] inj_on_imageI2[of g, OF inj_onI]) auto - then have "f x \ {min ?A ?B <..< max ?A ?B}" - by auto - ultimately - show ?thesis - by (simp add: continuous_on_eq_continuous_at) -qed - -lemma isCont_inverse_function2: - fixes f g :: "real \ real" shows - "\a < x; x < b; - \z. a \ z \ z \ b \ g (f z) = z; - \z. a \ z \ z \ b \ isCont f z\ - \ isCont g (f x)" -apply (rule isCont_inverse_function - [where f=f and d="min (x - a) (b - x)"]) -apply (simp_all add: abs_le_iff) -done - text {* Derivative of inverse function *} lemma DERIV_inverse_function: @@ -1228,44 +1014,6 @@ with g'cdef f'cdef cint show ?thesis by auto qed - -subsection {* Theorems about Limits *} - -(* need to rename second isCont_inverse *) - -lemma isCont_inv_fun: - fixes f g :: "real \ real" - shows "[| 0 < d; \z. \z - x\ \ d --> g(f(z)) = z; - \z. \z - x\ \ d --> isCont f z |] - ==> isCont g (f x)" -by (rule isCont_inverse_function) - -text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} -lemma LIM_fun_gt_zero: - "[| f -- c --> (l::real); 0 < l |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> 0 < f x)" -apply (drule (1) LIM_D, clarify) -apply (rule_tac x = s in exI) -apply (simp add: abs_less_iff) -done - -lemma LIM_fun_less_zero: - "[| f -- c --> (l::real); l < 0 |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x < 0)" -apply (drule LIM_D [where r="-l"], simp, clarify) -apply (rule_tac x = s in exI) -apply (simp add: abs_less_iff) -done - -lemma LIM_fun_not_zero: - "[| f -- c --> (l::real); l \ 0 |] - ==> \r. 0 < r & (\x::real. x \ c & \c - x\ < r --> f x \ 0)" -apply (rule linorder_cases [of l 0]) -apply (drule (1) LIM_fun_less_zero, force) -apply simp -apply (drule (1) LIM_fun_gt_zero, force) -done - lemma GMVT': fixes f g :: "real \ real" assumes "a < b" @@ -1284,6 +1032,9 @@ by auto qed + +subsection {* L'Hopitals rule *} + lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" unfolding DERIV_iff2