# HG changeset patch # User huffman # Date 1230154760 28800 # Node ID c3d1f87a3cb03d2eeef2130490d00bf0961b8203 # Parent 5eff800a695f1ac1f7aa26dbf50895d2271aa87d# Parent bad036eb71c4a364415c5ee159919c173cc5a394 merged. diff -r bad036eb71c4 -r c3d1f87a3cb0 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Tue Dec 23 21:24:40 2008 +0100 +++ b/src/HOL/Deriv.thy Wed Dec 24 13:39:20 2008 -0800 @@ -20,12 +20,6 @@ ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)" -definition - differentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" - (infixl "differentiable" 60) where - "f differentiable x = (\D. DERIV f x :> D)" - - consts Bolzano_bisect :: "[real*real=>bool, real, real, nat] => (real*real)" primrec @@ -316,63 +310,104 @@ subsection {* Differentiability predicate *} +definition + differentiable :: "['a::real_normed_field \ 'a, 'a] \ bool" + (infixl "differentiable" 60) where + "f differentiable x = (\D. DERIV f x :> D)" + +lemma differentiableE [elim?]: + assumes "f differentiable x" + obtains df where "DERIV f x :> df" + using prems unfolding differentiable_def .. + lemma differentiableD: "f differentiable x ==> \D. DERIV f x :> D" by (simp add: differentiable_def) lemma differentiableI: "DERIV f x :> D ==> f differentiable x" by (force simp add: differentiable_def) -lemma differentiable_const: "(\z. a) differentiable x" - apply (unfold differentiable_def) - apply (rule_tac x=0 in exI) - apply simp - done +lemma differentiable_ident [simp]: "(\x. x) differentiable x" + by (rule DERIV_ident [THEN differentiableI]) + +lemma differentiable_const [simp]: "(\z. a) differentiable x" + by (rule DERIV_const [THEN differentiableI]) -lemma differentiable_sum: +lemma differentiable_compose: + assumes f: "f differentiable (g x)" + assumes g: "g differentiable x" + shows "(\x. f (g x)) differentiable x" +proof - + from `f differentiable (g x)` obtain df where "DERIV f (g x) :> df" .. + moreover + from `g differentiable x` obtain dg where "DERIV g x :> dg" .. + ultimately + have "DERIV (\x. f (g x)) x :> df * dg" by (rule DERIV_chain2) + thus ?thesis by (rule differentiableI) +qed + +lemma differentiable_sum [simp]: assumes "f differentiable x" and "g differentiable x" shows "(\x. f x + g x) differentiable x" proof - - from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) - then obtain df where "DERIV f x :> df" .. - moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - ultimately have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) - hence "\D. DERIV (\x. f x + g x) x :> D" by auto - thus ?thesis by (fold differentiable_def) + from `f differentiable x` obtain df where "DERIV f x :> df" .. + moreover + from `g differentiable x` obtain dg where "DERIV g x :> dg" .. + ultimately + have "DERIV (\x. f x + g x) x :> df + dg" by (rule DERIV_add) + thus ?thesis by (rule differentiableI) +qed + +lemma differentiable_minus [simp]: + assumes "f differentiable x" + shows "(\x. - f x) differentiable x" +proof - + from `f differentiable x` obtain df where "DERIV f x :> df" .. + hence "DERIV (\x. - f x) x :> - df" by (rule DERIV_minus) + thus ?thesis by (rule differentiableI) qed -lemma differentiable_diff: +lemma differentiable_diff [simp]: assumes "f differentiable x" - and "g differentiable x" + assumes "g differentiable x" shows "(\x. f x - g x) differentiable x" + unfolding diff_minus using prems by simp + +lemma differentiable_mult [simp]: + assumes "f differentiable x" + assumes "g differentiable x" + shows "(\x. f x * g x) differentiable x" proof - - from prems have "f differentiable x" by simp + from `f differentiable x` obtain df where "DERIV f x :> df" .. moreover - from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - then have "DERIV (\x. - g x) x :> -dg" by (rule DERIV_minus) - hence "\D. DERIV (\x. - g x) x :> D" by auto - hence "(\x. - g x) differentiable x" by (fold differentiable_def) - ultimately - show ?thesis - by (auto simp: diff_def dest: differentiable_sum) + from `g differentiable x` obtain dg where "DERIV g x :> dg" .. + ultimately + have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (rule DERIV_mult) + thus ?thesis by (rule differentiableI) qed -lemma differentiable_mult: - assumes "f differentiable x" - and "g differentiable x" - shows "(\x. f x * g x) differentiable x" +lemma differentiable_inverse [simp]: + assumes "f differentiable x" and "f x \ 0" + shows "(\x. inverse (f x)) differentiable x" proof - - from prems have "\D. DERIV f x :> D" by (unfold differentiable_def) - then obtain df where "DERIV f x :> df" .. - moreover from prems have "\D. DERIV g x :> D" by (unfold differentiable_def) - then obtain dg where "DERIV g x :> dg" .. - ultimately have "DERIV (\x. f x * g x) x :> df * g x + dg * f x" by (simp add: DERIV_mult) - hence "\D. DERIV (\x. f x * g x) x :> D" by auto - thus ?thesis by (fold differentiable_def) + from `f differentiable x` obtain df where "DERIV f x :> df" .. + hence "DERIV (\x. inverse (f x)) x :> - (inverse (f x) * df * inverse (f x))" + using `f x \ 0` by (rule DERIV_inverse') + thus ?thesis by (rule differentiableI) qed +lemma differentiable_divide [simp]: + assumes "f differentiable x" + assumes "g differentiable x" and "g x \ 0" + shows "(\x. f x / g x) differentiable x" + unfolding divide_inverse using prems by simp + +lemma differentiable_power [simp]: + fixes f :: "'a::{recpower,real_normed_field} \ 'a" + assumes "f differentiable x" + shows "(\x. f x ^ n) differentiable x" + by (induct n, simp, simp add: power_Suc prems) + subsection {* Nested Intervals and Bisection *} @@ -1722,4 +1757,60 @@ apply (simp add: poly_entire del: pmult_Cons) done + +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) + +lemma isCont_inv_fun_inv: + fixes f g :: "real \ real" + shows "[| 0 < d; + \z. \z - x\ \ d --> g(f(z)) = z; + \z. \z - x\ \ d --> isCont f z |] + ==> \e. 0 < e & + (\y. 0 < \y - f(x)\ & \y - f(x)\ < e --> f(g(y)) = y)" +apply (drule isCont_inj_range) +prefer 2 apply (assumption, assumption, auto) +apply (rule_tac x = e in exI, auto) +apply (rotate_tac 2) +apply (drule_tac x = y in spec, auto) +done + + +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 (auto simp add: LIM_def) +apply (drule_tac x = "l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: 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 (auto simp add: LIM_def) +apply (drule_tac x = "-l/2" in spec, safe, force) +apply (rule_tac x = s in exI) +apply (auto simp only: 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 (cut_tac x = l and y = 0 in linorder_less_linear, auto) +apply (drule LIM_fun_less_zero) +apply (drule_tac [3] LIM_fun_gt_zero) +apply force+ +done + end diff -r bad036eb71c4 -r c3d1f87a3cb0 src/HOL/MacLaurin.thy --- a/src/HOL/MacLaurin.thy Tue Dec 23 21:24:40 2008 +0100 +++ b/src/HOL/MacLaurin.thy Wed Dec 24 13:39:20 2008 -0800 @@ -77,7 +77,7 @@ apply (rule_tac [2] DERIV_pow) prefer 3 apply (simp add: fact_diff_Suc) prefer 2 apply simp -apply (frule_tac m = m in less_add_one, clarify) +apply (frule less_iff_Suc_add [THEN iffD1], clarify) apply (simp del: setsum_op_ivl_Suc) apply (insert sumr_offset4 [of 1]) apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) @@ -150,7 +150,7 @@ prefer 2 apply clarify apply simp - apply (frule_tac m = ma in less_add_one, clarify) + apply (frule less_iff_Suc_add [THEN iffD1], clarify) apply (simp del: setsum_op_ivl_Suc) apply (insert sumr_offset4 [of 1]) apply (simp del: setsum_op_ivl_Suc fact_Suc realpow_Suc) diff -r bad036eb71c4 -r c3d1f87a3cb0 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Tue Dec 23 21:24:40 2008 +0100 +++ b/src/HOL/Transcendental.thy Wed Dec 24 13:39:20 2008 -0800 @@ -11,7 +11,7 @@ imports Fact Series Deriv NthRoot begin -subsection{*Properties of Power Series*} +subsection {* Properties of Power Series *} lemma lemma_realpow_diff: fixes y :: "'a::recpower" @@ -26,8 +26,8 @@ fixes y :: "'a::{recpower,comm_semiring_0}" shows "(\p=0..p=0.. 'a::ring_1) => nat => 'a" where @@ -124,33 +124,22 @@ lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)" by (simp add: diffs_def) -text{*Show that we can shift the terms down one*} -lemma lemma_diffs: - "(\n=0..n=0..n. f (Suc n)) sums s \ (\n. f n) sums s" +unfolding sums_def +apply (rule LIMSEQ_imp_Suc) +apply (subst setsum_shift_lb_Suc0_0_upt [where f=f, OF f, symmetric]) +apply (simp only: setsum_shift_bounds_Suc_ivl) done -lemma lemma_diffs2: - "(\n=0..n=0.. (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums (\n. (diffs c)(n) * (x ^ n))" -apply (subgoal_tac " (%n. of_nat n * c (n) * (x ^ (n - Suc 0))) ----> 0") -apply (rule_tac [2] LIMSEQ_imp_Suc) -apply (drule summable_sums) -apply (auto simp add: sums_def) -apply (drule_tac X="(\n. \n = 0.. (\d. n = m + d + Suc 0)" -by (simp add: less_iff_Suc_add) - -lemma sumdiff: "a + b - (c + d) = a - c + b - (d::real)" -by arith - lemma sumr_diff_mult_const2: "setsum f {0..i = 0..h. \h \ 0; norm h < k\ \ norm (f h) \ K * norm h" shows "f -- 0 --> 0" -proof (simp add: LIM_def, safe) +unfolding LIM_def diff_0_right +proof (safe) + let ?h = "of_real (k / 2)::'a" + have "?h \ 0" and "norm ?h < k" using k by simp_all + hence "norm (f ?h) \ K * norm ?h" by (rule le) + hence "0 \ K * norm ?h" by (rule order_trans [OF norm_ge_zero]) + hence zero_le_K: "0 \ K" using k by (simp add: zero_le_mult_iff) + fix r::real assume r: "0 < r" - have zero_le_K: "0 \ K" - apply (cut_tac k) - apply (cut_tac h="of_real (k/2)" in le, simp) - apply (simp del: of_real_divide) - apply (drule order_trans [OF norm_ge_zero]) - apply (simp add: zero_le_mult_iff) - done show "\s. 0 < s \ (\x. x \ 0 \ norm x < s \ norm (f x) < r)" proof (cases) assume "K = 0" @@ -392,11 +375,12 @@ assumes 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" assumes 4: "norm x < norm K" shows "DERIV (\x. \n. c n * x ^ n) x :> (\n. (diffs c) n * x ^ n)" -proof (simp add: deriv_def, rule LIM_zero_cancel) +unfolding deriv_def +proof (rule LIM_zero_cancel) show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x ^ n)) / h - suminf (\n. diffs c n * x ^ n)) -- 0 --> 0" proof (rule LIM_equal2) - show "0 < norm K - norm x" by (simp add: less_diff_eq 4) + show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq) next fix h :: 'a assume "h \ 0" @@ -421,8 +405,7 @@ apply (rule summable_divide) apply (rule summable_diff [OF B A]) apply (rule sums_summable [OF diffs_equiv [OF C]]) - apply (rule_tac f="suminf" in arg_cong) - apply (rule ext) + apply (rule arg_cong [where f="suminf"], rule ext) apply (simp add: ring_simps) done next @@ -433,22 +416,12 @@ qed -subsection{*Exponential Function*} +subsection {* Exponential Function *} definition exp :: "'a \ 'a::{recpower,real_normed_field,banach}" where "exp x = (\n. x ^ n /\<^sub>R real (fact n))" -definition - sin :: "real => real" where - "sin x = (\n. (if even(n) then 0 else - (-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)" - -definition - cos :: "real => real" where - "cos x = (\n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) - else 0) * x ^ n)" - lemma summable_exp_generic: fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" defines S_def: "S \ \n. x ^ n /\<^sub>R real (fact n)" @@ -493,66 +466,9 @@ lemma summable_exp: "summable (%n. inverse (real (fact n)) * x ^ n)" by (insert summable_exp_generic [where x=x], simp) -lemma summable_sin: - "summable (%n. - (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n)" -apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) -apply (rule_tac [2] summable_exp) -apply (rule_tac x = 0 in exI) -apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) -done - -lemma summable_cos: - "summable (%n. - (if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" -apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) -apply (rule_tac [2] summable_exp) -apply (rule_tac x = 0 in exI) -apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) -done - -lemma lemma_STAR_sin: - "(if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos: - "0 < n --> - -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos1: - "0 < n --> - (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" -by (induct "n", auto) - -lemma lemma_STAR_cos2: - "(\n=1..n. x ^ n /\<^sub>R real (fact n)) sums exp x" unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) -lemma sin_converges: - "(%n. (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n) sums sin(x)" -unfolding sin_def by (rule summable_sin [THEN summable_sums]) - -lemma cos_converges: - "(%n. (if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0) * x ^ n) sums cos(x)" -unfolding cos_def by (rule summable_cos [THEN summable_sums]) - - -subsection{*Formal Derivatives of Exp, Sin, and Cos Series*} lemma exp_fdiffs: "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))" @@ -562,48 +478,6 @@ lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" by (simp add: diffs_def) -lemma sin_fdiffs: - "diffs(%n. if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) - = (%n. if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0)" -by (auto intro!: ext - simp add: diffs_def divide_inverse real_of_nat_def of_nat_mult - simp del: mult_Suc of_nat_Suc) - -lemma sin_fdiffs2: - "diffs(%n. if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) n - = (if even n then - -1 ^ (n div 2)/(real (fact n)) - else 0)" -by (simp only: sin_fdiffs) - -lemma cos_fdiffs: - "diffs(%n. if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) - = (%n. - (if even n then 0 - else -1 ^ ((n - Suc 0)div 2)/(real (fact n))))" -by (auto intro!: ext - simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def of_nat_mult - simp del: mult_Suc of_nat_Suc) - - -lemma cos_fdiffs2: - "diffs(%n. if even n then - -1 ^ (n div 2)/(real (fact n)) else 0) n - = - (if even n then 0 - else -1 ^ ((n - Suc 0)div 2)/(real (fact n)))" -by (simp only: cos_fdiffs) - -text{*Now at last we can get the derivatives of exp, sin and cos*} - -lemma lemma_sin_minus: - "- sin x = (\n. - ((if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" -by (auto intro!: sums_unique sums_minus sin_converges) - lemma lemma_exp_ext: "exp = (\x. \n. x ^ n /\<^sub>R real (fact n))" by (auto intro!: ext simp add: exp_def) @@ -617,45 +491,11 @@ apply (simp del: of_real_add) done -lemma lemma_sin_ext: - "sin = (%x. \n. - (if even n then 0 - else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * - x ^ n)" -by (auto intro!: ext simp add: sin_def) - -lemma lemma_cos_ext: - "cos = (%x. \n. - (if even n then -1 ^ (n div 2)/(real (fact n)) else 0) * - x ^ n)" -by (auto intro!: ext simp add: cos_def) - -lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" -apply (simp add: cos_def) -apply (subst lemma_sin_ext) -apply (auto simp add: sin_fdiffs2 [symmetric]) -apply (rule_tac K = "1 + \x\" in termdiffs) -apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs) -done - -lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" -apply (subst lemma_cos_ext) -apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) -apply (rule_tac K = "1 + \x\" in termdiffs) -apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus) -done - lemma isCont_exp [simp]: "isCont exp x" by (rule DERIV_exp [THEN DERIV_isCont]) -lemma isCont_sin [simp]: "isCont sin x" -by (rule DERIV_sin [THEN DERIV_isCont]) -lemma isCont_cos [simp]: "isCont cos x" -by (rule DERIV_cos [THEN DERIV_isCont]) - - -subsection{*Properties of the Exponential Function*} +subsubsection {* Properties of the Exponential Function *} lemma powser_zero: fixes f :: "nat \ 'a::{real_normed_algebra_1,recpower}" @@ -724,6 +564,9 @@ unfolding exp_def by (simp only: Cauchy_product summable_norm_exp exp_series_add) +lemma mult_exp_exp: "exp x * exp y = exp (x + y)" +by (rule exp_add [symmetric]) + lemma exp_of_real: "exp (of_real x) = of_real (exp x)" unfolding exp_def apply (subst of_real.suminf) @@ -731,6 +574,51 @@ apply (simp add: scaleR_conv_of_real) done +lemma exp_not_eq_zero [simp]: "exp x \ 0" +proof + have "exp x * exp (- x) = 1" by (simp add: mult_exp_exp) + also assume "exp x = 0" + finally show "False" by simp +qed + +lemma exp_minus: "exp (- x) = inverse (exp x)" +by (rule inverse_unique [symmetric], simp add: mult_exp_exp) + +lemma exp_diff: "exp (x - y) = exp x / exp y" + unfolding diff_minus divide_inverse + by (simp add: exp_add exp_minus) + + +subsubsection {* Properties of the Exponential Function on Reals *} + +text {* Comparisons of @{term "exp x"} with zero. *} + +text{*Proof: because every exponential can be seen as a square.*} +lemma exp_ge_zero [simp]: "0 \ exp (x::real)" +proof - + have "0 \ exp (x/2) * exp (x/2)" by simp + thus ?thesis by (simp add: exp_add [symmetric]) +qed + +lemma exp_gt_zero [simp]: "0 < exp (x::real)" +by (simp add: order_less_le) + +lemma not_exp_less_zero [simp]: "\ exp (x::real) < 0" +by (simp add: not_less) + +lemma not_exp_le_zero [simp]: "\ exp (x::real) \ 0" +by (simp add: not_le) + +lemma abs_exp_cancel [simp]: "\exp x::real\ = exp x" +by simp + +lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" +apply (induct "n") +apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) +done + +text {* Strict monotonicity of exponential. *} + lemma exp_ge_add_one_self_aux: "0 \ (x::real) ==> (1 + x) \ exp(x)" apply (drule order_le_imp_less_or_eq, auto) apply (simp add: exp_def) @@ -739,114 +627,61 @@ apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_mult_iff) done -lemma exp_gt_one [simp]: "0 < (x::real) ==> 1 < exp x" -apply (rule order_less_le_trans) -apply (rule_tac [2] exp_ge_add_one_self_aux, auto) -done - -lemma DERIV_exp_add_const: "DERIV (%x. exp (x + y)) x :> exp(x + y)" -proof - - have "DERIV (exp \ (\x. x + y)) x :> exp (x + y) * (1+0)" - by (fast intro: DERIV_chain DERIV_add DERIV_exp DERIV_ident DERIV_const) - thus ?thesis by (simp add: o_def) -qed - -lemma DERIV_exp_minus [simp]: "DERIV (%x. exp (-x)) x :> - exp(-x)" +lemma exp_gt_one: "0 < (x::real) \ 1 < exp x" proof - - have "DERIV (exp \ uminus) x :> exp (- x) * - 1" - by (fast intro: DERIV_chain DERIV_minus DERIV_exp DERIV_ident) - thus ?thesis by (simp add: o_def) -qed - -lemma DERIV_exp_exp_zero [simp]: "DERIV (%x. exp (x + y) * exp (- x)) x :> 0" -proof - - have "DERIV (\x. exp (x + y) * exp (- x)) x - :> exp (x + y) * exp (- x) + - exp (- x) * exp (x + y)" - by (fast intro: DERIV_exp_add_const DERIV_exp_minus DERIV_mult) - thus ?thesis by (simp add: mult_commute) -qed - -lemma exp_add_mult_minus [simp]: "exp(x + y)*exp(-x) = exp(y::real)" -proof - - have "\x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0" by simp - hence "exp (x + y) * exp (- x) = exp (0 + y) * exp (- 0)" - by (rule DERIV_isconst_all) - thus ?thesis by simp + assume x: "0 < x" + hence "1 < 1 + x" by simp + also from x have "1 + x \ exp x" + by (simp add: exp_ge_add_one_self_aux) + finally show ?thesis . qed -lemma exp_mult_minus [simp]: "exp x * exp(-x) = 1" -by (simp add: exp_add [symmetric]) - -lemma exp_mult_minus2 [simp]: "exp(-x)*exp(x) = 1" -by (simp add: mult_commute) - - -lemma exp_minus: "exp(-x) = inverse(exp(x))" -by (auto intro: inverse_unique [symmetric]) - -text{*Proof: because every exponential can be seen as a square.*} -lemma exp_ge_zero [simp]: "0 \ exp (x::real)" -apply (rule_tac t = x in real_sum_of_halves [THEN subst]) -apply (subst exp_add, auto) -done - -lemma exp_not_eq_zero [simp]: "exp x \ 0" -apply (cut_tac x = x in exp_mult_minus2) -apply (auto simp del: exp_mult_minus2) -done - -lemma exp_gt_zero [simp]: "0 < exp (x::real)" -by (simp add: order_less_le) - -lemma inv_exp_gt_zero [simp]: "0 < inverse(exp x::real)" -by (auto intro: positive_imp_inverse_positive) - -lemma abs_exp_cancel [simp]: "\exp x::real\ = exp x" -by auto - -lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" -apply (induct "n") -apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) -done - -lemma exp_diff: "exp(x - y) = exp(x)/(exp y)" -apply (simp add: diff_minus divide_inverse) -apply (simp (no_asm) add: exp_add exp_minus) -done - - lemma exp_less_mono: fixes x y :: real - assumes xy: "x < y" shows "exp x < exp y" + assumes "x < y" shows "exp x < exp y" proof - - from xy have "1 < exp (y + - x)" - by (rule real_less_sum_gt_zero [THEN exp_gt_one]) - hence "exp x * inverse (exp x) < exp y * inverse (exp x)" - by (auto simp add: exp_add exp_minus) - thus ?thesis - by (simp add: divide_inverse [symmetric] pos_less_divide_eq - del: divide_self_if) + from `x < y` have "0 < y - x" by simp + hence "1 < exp (y - x)" by (rule exp_gt_one) + hence "1 < exp y / exp x" by (simp only: exp_diff) + thus "exp x < exp y" by simp qed lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y" -apply (simp add: linorder_not_le [symmetric]) -apply (auto simp add: order_le_less exp_less_mono) +apply (simp add: linorder_not_le [symmetric]) +apply (auto simp add: order_le_less exp_less_mono) done -lemma exp_less_cancel_iff [iff]: "(exp(x::real) < exp(y)) = (x < y)" +lemma exp_less_cancel_iff [iff]: "exp (x::real) < exp y \ x < y" by (auto intro: exp_less_mono exp_less_cancel) -lemma exp_le_cancel_iff [iff]: "(exp(x::real) \ exp(y)) = (x \ y)" +lemma exp_le_cancel_iff [iff]: "exp (x::real) \ exp y \ x \ y" by (auto simp add: linorder_not_less [symmetric]) -lemma exp_inj_iff [iff]: "(exp (x::real) = exp y) = (x = y)" +lemma exp_inj_iff [iff]: "exp (x::real) = exp y \ x = y" by (simp add: order_eq_iff) +text {* Comparisons of @{term "exp x"} with one. *} + +lemma one_less_exp_iff [simp]: "1 < exp (x::real) \ 0 < x" + using exp_less_cancel_iff [where x=0 and y=x] by simp + +lemma exp_less_one_iff [simp]: "exp (x::real) < 1 \ x < 0" + using exp_less_cancel_iff [where x=x and y=0] by simp + +lemma one_le_exp_iff [simp]: "1 \ exp (x::real) \ 0 \ x" + using exp_le_cancel_iff [where x=0 and y=x] by simp + +lemma exp_le_one_iff [simp]: "exp (x::real) \ 1 \ x \ 0" + using exp_le_cancel_iff [where x=x and y=0] by simp + +lemma exp_eq_one_iff [simp]: "exp (x::real) = 1 \ x = 0" + using exp_inj_iff [where x=x and y=0] by simp + lemma lemma_exp_total: "1 \ y ==> \x. 0 \ x & x \ y - 1 & exp(x::real) = y" apply (rule IVT) apply (auto intro: isCont_exp simp add: le_diff_eq) apply (subgoal_tac "1 + (y - 1) \ exp (y - 1)") -apply simp +apply simp apply (rule exp_ge_add_one_self_aux, simp) done @@ -861,7 +696,7 @@ done -subsection{*Properties of the Logarithmic Function*} +subsection {* Natural Logarithm *} definition ln :: "real => real" where @@ -873,59 +708,46 @@ lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" by (auto dest: exp_total) -lemma exp_ln_iff [simp]: "(exp (ln x) = x) = (0 < x)" -apply (auto dest: exp_total) -apply (erule subst, simp) -done - -lemma ln_mult: "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)" -apply (rule exp_inj_iff [THEN iffD1]) -apply (simp add: exp_add exp_ln mult_pos_pos) -done - -lemma ln_inj_iff[simp]: "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)" -apply (simp only: exp_ln_iff [symmetric]) -apply (erule subst)+ -apply simp -done - -lemma ln_one[simp]: "ln 1 = 0" -by (rule exp_inj_iff [THEN iffD1], auto) - -lemma ln_inverse: "0 < x ==> ln(inverse x) = - ln x" -apply (rule_tac a1 = "ln x" in add_left_cancel [THEN iffD1]) -apply (auto simp add: positive_imp_inverse_positive ln_mult [symmetric]) +lemma exp_ln_iff [simp]: "exp (ln x) = x \ 0 < x" +apply (rule iffI) +apply (erule subst, rule exp_gt_zero) +apply (erule exp_ln) done -lemma ln_div: - "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y" -apply (simp add: divide_inverse) -apply (auto simp add: positive_imp_inverse_positive ln_mult ln_inverse) -done +lemma ln_unique: "exp y = x \ ln x = y" +by (erule subst, rule ln_exp) + +lemma ln_one [simp]: "ln 1 = 0" +by (rule ln_unique, simp) + +lemma ln_mult: "\0 < x; 0 < y\ \ ln (x * y) = ln x + ln y" +by (rule ln_unique, simp add: exp_add) + +lemma ln_inverse: "0 < x \ ln (inverse x) = - ln x" +by (rule ln_unique, simp add: exp_minus) + +lemma ln_div: "\0 < x; 0 < y\ \ ln (x / y) = ln x - ln y" +by (rule ln_unique, simp add: exp_diff) -lemma ln_less_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)" -apply (simp only: exp_ln_iff [symmetric]) -apply (erule subst)+ -apply simp +lemma ln_realpow: "0 < x \ ln (x ^ n) = real n * ln x" +by (rule ln_unique, simp add: exp_real_of_nat_mult) + +lemma ln_less_cancel_iff [simp]: "\0 < x; 0 < y\ \ ln x < ln y \ x < y" +by (subst exp_less_cancel_iff [symmetric], simp) + +lemma ln_le_cancel_iff [simp]: "\0 < x; 0 < y\ \ ln x \ ln y \ x \ y" +by (simp add: linorder_not_less [symmetric]) + +lemma ln_inj_iff [simp]: "\0 < x; 0 < y\ \ ln x = ln y \ x = y" +by (simp add: order_eq_iff) + +lemma ln_add_one_self_le_self [simp]: "0 \ x \ ln (1 + x) \ x" +apply (rule exp_le_cancel_iff [THEN iffD1]) +apply (simp add: exp_ge_add_one_self_aux) done -lemma ln_le_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x \ ln y) = (x \ y)" -by (auto simp add: linorder_not_less [symmetric]) - -lemma ln_realpow: "0 < x ==> ln(x ^ n) = real n * ln(x)" -by (auto dest!: exp_total simp add: exp_real_of_nat_mult [symmetric]) - -lemma ln_add_one_self_le_self [simp]: "0 \ x ==> ln(1 + x) \ x" -apply (rule ln_exp [THEN subst]) -apply (rule ln_le_cancel_iff [THEN iffD2]) -apply (auto simp add: exp_ge_add_one_self_aux) -done - -lemma ln_less_self [simp]: "0 < x ==> ln x < x" -apply (rule order_less_le_trans) -apply (rule_tac [2] ln_add_one_self_le_self) -apply (rule ln_less_cancel_iff [THEN iffD2], auto) -done +lemma ln_less_self [simp]: "0 < x \ ln x < x" +by (rule order_less_le_trans [where y="ln (1 + x)"]) simp_all lemma ln_ge_zero [simp]: assumes x: "1 \ x" shows "0 \ ln x" @@ -992,7 +814,151 @@ done -subsection{*Basic Properties of the Trigonometric Functions*} +subsection {* Sine and Cosine *} + +definition + sin :: "real => real" where + "sin x = (\n. (if even(n) then 0 else + (-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)" + +definition + cos :: "real => real" where + "cos x = (\n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) + else 0) * x ^ n)" + +lemma summable_sin: + "summable (%n. + (if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n)" +apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +done + +lemma summable_cos: + "summable (%n. + (if even n then + -1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" +apply (rule_tac g = "(%n. inverse (real (fact n)) * \x\ ^ n)" in summable_comparison_test) +apply (rule_tac [2] summable_exp) +apply (rule_tac x = 0 in exI) +apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) +done + +lemma lemma_STAR_sin: + "(if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos: + "0 < n --> + -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos1: + "0 < n --> + (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" +by (induct "n", auto) + +lemma lemma_STAR_cos2: + "(\n=1..n. - ((if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" +by (auto intro!: sums_unique sums_minus sin_converges) + +lemma lemma_sin_ext: + "sin = (%x. \n. + (if even n then 0 + else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * + x ^ n)" +by (auto intro!: ext simp add: sin_def) + +lemma lemma_cos_ext: + "cos = (%x. \n. + (if even n then -1 ^ (n div 2)/(real (fact n)) else 0) * + x ^ n)" +by (auto intro!: ext simp add: cos_def) + +lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" +apply (simp add: cos_def) +apply (subst lemma_sin_ext) +apply (auto simp add: sin_fdiffs2 [symmetric]) +apply (rule_tac K = "1 + \x\" in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs) +done + +lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" +apply (subst lemma_cos_ext) +apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) +apply (rule_tac K = "1 + \x\" in termdiffs) +apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus) +done + +lemma isCont_sin [simp]: "isCont sin x" +by (rule DERIV_sin [THEN DERIV_isCont]) + +lemma isCont_cos [simp]: "isCont cos x" +by (rule DERIV_cos [THEN DERIV_isCont]) + + +subsection {* Properties of Sine and Cosine *} lemma sin_zero [simp]: "sin 0 = 0" unfolding sin_def by (simp add: powser_zero) @@ -1088,9 +1054,6 @@ apply (simp del: realpow_Suc) done -lemma real_gt_one_ge_zero_add_less: "[| 1 < x; 0 \ y |] ==> 1 < x + (y::real)" -by arith - lemma abs_sin_le_one [simp]: "\sin x\ \ 1" by (rule power2_le_imp_le, simp_all add: sin_squared_eq) @@ -1187,7 +1150,7 @@ apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) done -lemma sin_cos_minus [simp]: +lemma sin_cos_minus: "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0" apply (cut_tac y = 0 and x = x in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all]) @@ -1195,14 +1158,10 @@ done lemma sin_minus [simp]: "sin (-x) = -sin(x)" -apply (cut_tac x = x in sin_cos_minus) -apply (simp del: sin_cos_minus) -done + using sin_cos_minus [where x=x] by simp lemma cos_minus [simp]: "cos (-x) = cos(x)" -apply (cut_tac x = x in sin_cos_minus) -apply (simp del: sin_cos_minus) -done + using sin_cos_minus [where x=x] by simp lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" by (simp add: diff_minus sin_add) @@ -1217,16 +1176,14 @@ by (simp add: cos_diff mult_commute) lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" -by (cut_tac x = x and y = x in sin_add, auto) - + using sin_add [where x=x and y=x] by simp lemma cos_double: "cos(2* x) = ((cos x)\) - ((sin x)\)" -apply (cut_tac x = x and y = x in cos_add) -apply (simp add: power2_eq_square) -done + using cos_add [where x=x and y=x] + by (simp add: power2_eq_square) -subsection{*The Constant Pi*} +subsection {* The Constant Pi *} definition pi :: "real" where @@ -1401,8 +1358,8 @@ lemma pi_not_less_zero [simp]: "\ pi < 0" by (simp add: linorder_not_less) -lemma minus_pi_half_less_zero [simp]: "-(pi/2) < 0" -by auto +lemma minus_pi_half_less_zero: "-(pi/2) < 0" +by simp lemma sin_pi_half [simp]: "sin(pi/2) = 1" apply (cut_tac x = "pi/2" in sin_cos_squared_add2) @@ -1614,7 +1571,7 @@ done -subsection{*Tangent*} +subsection {* Tangent *} definition tan :: "real => real" where @@ -2139,11 +2096,6 @@ lemma sin_zero_abs_cos_one: "sin x = 0 ==> \cos x\ = 1" by (auto simp add: sin_zero_iff even_mult_two_ex) -lemma exp_eq_one_iff [simp]: "(exp (x::real) = 1) = (x = 0)" -apply auto -apply (drule_tac f = ln in arg_cong, auto) -done - lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0" by (cut_tac x = x in sin_cos_squared_add3, auto) @@ -2190,60 +2142,4 @@ apply (erule polar_ex2) done - -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) - -lemma isCont_inv_fun_inv: - fixes f g :: "real \ real" - shows "[| 0 < d; - \z. \z - x\ \ d --> g(f(z)) = z; - \z. \z - x\ \ d --> isCont f z |] - ==> \e. 0 < e & - (\y. 0 < \y - f(x)\ & \y - f(x)\ < e --> f(g(y)) = y)" -apply (drule isCont_inj_range) -prefer 2 apply (assumption, assumption, auto) -apply (rule_tac x = e in exI, auto) -apply (rotate_tac 2) -apply (drule_tac x = y in spec, auto) -done - - -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 (auto simp add: LIM_def) -apply (drule_tac x = "l/2" in spec, safe, force) -apply (rule_tac x = s in exI) -apply (auto simp only: 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 (auto simp add: LIM_def) -apply (drule_tac x = "-l/2" in spec, safe, force) -apply (rule_tac x = s in exI) -apply (auto simp only: 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 (cut_tac x = l and y = 0 in linorder_less_linear, auto) -apply (drule LIM_fun_less_zero) -apply (drule_tac [3] LIM_fun_gt_zero) -apply force+ -done - end