# HG changeset patch # User haftmann # Date 1313825423 -7200 # Node ID 46d5e7f52ba5f118d90820449642a86b4e9d8264 # Parent 33439faadd6715bb2b06bbc93de2916c3070f562# Parent 2b088d74beb3dc949a87e1413cc4fb302f086b89 merged diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/Complex.thy --- a/src/HOL/Complex.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/Complex.thy Sat Aug 20 09:30:23 2011 +0200 @@ -606,14 +606,6 @@ abbreviation expi :: "complex \ complex" where "expi \ exp" -lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" - unfolding cos_coeff_def sin_coeff_def - by (simp del: mult_Suc, auto simp add: odd_Suc_mult_two_ex) - -lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" - unfolding cos_coeff_def sin_coeff_def - by (simp del: mult_Suc) - lemma expi_imaginary: "expi (Complex 0 b) = cis b" proof (rule complex_eqI) { fix n have "Complex 0 b ^ n = diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/Deriv.thy Sat Aug 20 09:30:23 2011 +0200 @@ -304,9 +304,6 @@ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))" by (drule (2) DERIV_divide) (simp add: mult_commute) -lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" -by auto - text {* @{text "DERIV_intros"} *} ML {* structure Deriv_Intros = Named_Thms @@ -524,7 +521,7 @@ ((\n. l \ g(n)) & g ----> l)" apply (drule lemma_nest, auto) apply (subgoal_tac "l = m") -apply (drule_tac [2] X = f in LIMSEQ_diff) +apply (drule_tac [2] f = f in LIMSEQ_diff) apply (auto intro: LIMSEQ_unique) done diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/Library/Poly_Deriv.thy --- a/src/HOL/Library/Poly_Deriv.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/Library/Poly_Deriv.thy Sat Aug 20 09:30:23 2011 +0200 @@ -78,11 +78,11 @@ by (simp add: DERIV_cmult mult_commute [of _ c]) lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" -by (rule lemma_DERIV_subst, rule DERIV_pow, simp) +by (rule DERIV_cong, rule DERIV_pow, simp) declare DERIV_pow2 [simp] DERIV_pow [simp] lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D" -by (rule lemma_DERIV_subst, rule DERIV_add, auto) +by (rule DERIV_cong, rule DERIV_add, auto) lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x" by (induct p, auto intro!: DERIV_intros simp add: pderiv_pCons) diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/Lim.thy --- a/src/HOL/Lim.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/Lim.thy Sat Aug 20 09:30:23 2011 +0200 @@ -81,38 +81,8 @@ shows "(\h. f (a + h)) -- 0 --> L \ f -- a --> L" by (drule_tac k="- a" in LIM_offset, simp) -lemma LIM_const [simp]: "(%x. k) -- x --> k" -by (rule tendsto_const) - lemma LIM_cong_limit: "\ f -- x --> L ; K = L \ \ f -- x --> K" by simp -lemma LIM_add: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - assumes f: "f -- a --> L" and g: "g -- a --> M" - shows "(\x. f x + g x) -- a --> (L + M)" -using assms by (rule tendsto_add) - -lemma LIM_add_zero: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - shows "\f -- a --> 0; g -- a --> 0\ \ (\x. f x + g x) -- a --> 0" - by (rule tendsto_add_zero) - -lemma LIM_minus: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "f -- a --> L \ (\x. - f x) -- a --> - L" -by (rule tendsto_minus) - -(* TODO: delete *) -lemma LIM_add_minus: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - shows "[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + -g(x)) -- x --> (l + -m)" -by (intro LIM_add LIM_minus) - -lemma LIM_diff: - fixes f g :: "'a::topological_space \ 'b::real_normed_vector" - shows "\f -- x --> l; g -- x --> m\ \ (\x. f x - g x) -- x --> l - m" -by (rule tendsto_diff) - lemma LIM_zero: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "f -- a --> l \ (\x. f x - l) -- a --> 0" @@ -144,38 +114,6 @@ by (rule metric_LIM_imp_LIM [OF f], simp add: dist_norm le) -lemma LIM_norm: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "f -- a --> l \ (\x. norm (f x)) -- a --> norm l" -by (rule tendsto_norm) - -lemma LIM_norm_zero: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "f -- a --> 0 \ (\x. norm (f x)) -- a --> 0" -by (rule tendsto_norm_zero) - -lemma LIM_norm_zero_cancel: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "(\x. norm (f x)) -- a --> 0 \ f -- a --> 0" -by (rule tendsto_norm_zero_cancel) - -lemma LIM_norm_zero_iff: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "(\x. norm (f x)) -- a --> 0 = f -- a --> 0" -by (rule tendsto_norm_zero_iff) - -lemma LIM_rabs: "f -- a --> (l::real) \ (\x. \f x\) -- a --> \l\" - by (rule tendsto_rabs) - -lemma LIM_rabs_zero: "f -- a --> (0::real) \ (\x. \f x\) -- a --> 0" - by (rule tendsto_rabs_zero) - -lemma LIM_rabs_zero_cancel: "(\x. \f x\) -- a --> (0::real) \ f -- a --> 0" - by (rule tendsto_rabs_zero_cancel) - -lemma LIM_rabs_zero_iff: "(\x. \f x\) -- a --> (0::real) = f -- a --> 0" - by (rule tendsto_rabs_zero_iff) - lemma trivial_limit_at: fixes a :: "'a::real_normed_algebra_1" shows "\ trivial_limit (at a)" -- {* TODO: find a more appropriate class *} @@ -203,9 +141,6 @@ shows "\f -- a --> L; f -- a --> M\ \ L = M" using trivial_limit_at by (rule tendsto_unique) -lemma LIM_ident [simp]: "(\x. x) -- a --> a" -by (rule tendsto_ident_at) - text{*Limits are equal for functions equal except at limit point*} lemma LIM_equal: "[| \x. x \ a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)" @@ -235,20 +170,6 @@ shows "g -- a --> l \ f -- a --> l" by (rule metric_LIM_equal2 [OF 1 2], simp_all add: dist_norm) -text{*Two uses in Transcendental.ML*} (* BH: no longer true; delete? *) -lemma LIM_trans: - fixes f g :: "'a::metric_space \ 'b::real_normed_vector" - shows "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] ==> f -- a --> l" -apply (drule LIM_add, assumption) -apply (auto simp add: add_assoc) -done - -lemma LIM_compose: - assumes g: "g -- l --> g l" - assumes f: "f -- a --> l" - shows "(\x. g (f x)) -- a --> g l" - using assms by (rule tendsto_compose) - lemma LIM_compose_eventually: assumes f: "f -- a --> b" assumes g: "g -- b --> c" @@ -261,8 +182,8 @@ assumes g: "g -- b --> c" assumes inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" shows "(\x. g (f x)) -- a --> c" -using f g inj [folded eventually_at] -by (rule LIM_compose_eventually) + using g f inj [folded eventually_at] + by (rule tendsto_compose_eventually) lemma LIM_compose2: fixes a :: "'a::real_normed_vector" @@ -273,7 +194,7 @@ by (rule metric_LIM_compose2 [OF f g inj [folded dist_norm]]) lemma LIM_o: "\g -- l --> g l; f -- a --> l\ \ (g \ f) -- a --> g l" -unfolding o_def by (rule LIM_compose) + unfolding o_def by (rule tendsto_compose) lemma real_LIM_sandwich_zero: fixes f g :: "'a::topological_space \ real" @@ -321,9 +242,6 @@ "f -- a --> 0 \ (\x. c ** f x) -- a --> 0" by (rule tendsto_right_zero) -lemmas LIM_mult = - bounded_bilinear.LIM [OF bounded_bilinear_mult] - lemmas LIM_mult_zero = bounded_bilinear.LIM_prod_zero [OF bounded_bilinear_mult] @@ -333,32 +251,10 @@ lemmas LIM_mult_right_zero = bounded_bilinear.LIM_right_zero [OF bounded_bilinear_mult] -lemmas LIM_scaleR = - bounded_bilinear.LIM [OF bounded_bilinear_scaleR] - -lemmas LIM_of_real = - bounded_linear.LIM [OF bounded_linear_of_real] - -lemma LIM_power: - fixes f :: "'a::topological_space \ 'b::{power,real_normed_algebra}" - assumes f: "f -- a --> l" - shows "(\x. f x ^ n) -- a --> l ^ n" - using assms by (rule tendsto_power) - -lemma LIM_inverse: - fixes L :: "'a::real_normed_div_algebra" - shows "\f -- a --> L; L \ 0\ \ (\x. inverse (f x)) -- a --> inverse L" -by (rule tendsto_inverse) - lemma LIM_inverse_fun: assumes a: "a \ (0::'a::real_normed_div_algebra)" shows "inverse -- a --> inverse a" -by (rule LIM_inverse [OF LIM_ident a]) - -lemma LIM_sgn: - fixes f :: "'a::topological_space \ 'b::real_normed_vector" - shows "\f -- a --> l; l \ 0\ \ (\x. sgn (f x)) -- a --> sgn l" - by (rule tendsto_sgn) + by (rule tendsto_inverse [OF tendsto_ident_at a]) subsection {* Continuity *} @@ -374,54 +270,54 @@ by (simp add: isCont_def LIM_isCont_iff) lemma isCont_ident [simp]: "isCont (\x. x) a" - unfolding isCont_def by (rule LIM_ident) + unfolding isCont_def by (rule tendsto_ident_at) lemma isCont_const [simp]: "isCont (\x. k) a" - unfolding isCont_def by (rule LIM_const) + unfolding isCont_def by (rule tendsto_const) lemma isCont_norm [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "isCont f a \ isCont (\x. norm (f x)) a" - unfolding isCont_def by (rule LIM_norm) + unfolding isCont_def by (rule tendsto_norm) lemma isCont_rabs [simp]: fixes f :: "'a::topological_space \ real" shows "isCont f a \ isCont (\x. \f x\) a" - unfolding isCont_def by (rule LIM_rabs) + unfolding isCont_def by (rule tendsto_rabs) lemma isCont_add [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "\isCont f a; isCont g a\ \ isCont (\x. f x + g x) a" - unfolding isCont_def by (rule LIM_add) + unfolding isCont_def by (rule tendsto_add) lemma isCont_minus [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "isCont f a \ isCont (\x. - f x) a" - unfolding isCont_def by (rule LIM_minus) + unfolding isCont_def by (rule tendsto_minus) lemma isCont_diff [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "\isCont f a; isCont g a\ \ isCont (\x. f x - g x) a" - unfolding isCont_def by (rule LIM_diff) + unfolding isCont_def by (rule tendsto_diff) lemma isCont_mult [simp]: fixes f g :: "'a::topological_space \ 'b::real_normed_algebra" shows "\isCont f a; isCont g a\ \ isCont (\x. f x * g x) a" - unfolding isCont_def by (rule LIM_mult) + unfolding isCont_def by (rule tendsto_mult) lemma isCont_inverse [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_div_algebra" shows "\isCont f a; f a \ 0\ \ isCont (\x. inverse (f x)) a" - unfolding isCont_def by (rule LIM_inverse) + unfolding isCont_def by (rule tendsto_inverse) lemma isCont_divide [simp]: fixes f g :: "'a::topological_space \ 'b::real_normed_field" shows "\isCont f a; isCont g a; g a \ 0\ \ isCont (\x. f x / g x) a" unfolding isCont_def by (rule tendsto_divide) -lemma isCont_LIM_compose: - "\isCont g l; f -- a --> l\ \ (\x. g (f x)) -- a --> g l" - unfolding isCont_def by (rule LIM_compose) +lemma isCont_tendsto_compose: + "\isCont g l; (f ---> l) F\ \ ((\x. g (f x)) ---> g l) F" + unfolding isCont_def by (rule tendsto_compose) lemma metric_isCont_LIM_compose2: assumes f [unfolded isCont_def]: "isCont f a" @@ -439,18 +335,18 @@ by (rule LIM_compose2 [OF f g inj]) lemma isCont_o2: "\isCont f a; isCont g (f a)\ \ isCont (\x. g (f x)) a" - unfolding isCont_def by (rule LIM_compose) + unfolding isCont_def by (rule tendsto_compose) lemma isCont_o: "\isCont f a; isCont g (f a)\ \ isCont (g o f) a" unfolding o_def by (rule isCont_o2) lemma (in bounded_linear) isCont: "isCont g a \ isCont (\x. f (g x)) a" - unfolding isCont_def by (rule LIM) + unfolding isCont_def by (rule tendsto) lemma (in bounded_bilinear) isCont: "\isCont f a; isCont g a\ \ isCont (\x. f x ** g x) a" - unfolding isCont_def by (rule LIM) + unfolding isCont_def by (rule tendsto) lemmas isCont_scaleR [simp] = bounded_bilinear.isCont [OF bounded_bilinear_scaleR] @@ -461,12 +357,12 @@ lemma isCont_power [simp]: fixes f :: "'a::topological_space \ 'b::{power,real_normed_algebra}" shows "isCont f a \ isCont (\x. f x ^ n) a" - unfolding isCont_def by (rule LIM_power) + unfolding isCont_def by (rule tendsto_power) lemma isCont_sgn [simp]: fixes f :: "'a::topological_space \ 'b::real_normed_vector" shows "\isCont f a; f a \ 0\ \ isCont (\x. sgn (f x)) a" - unfolding isCont_def by (rule LIM_sgn) + unfolding isCont_def by (rule tendsto_sgn) lemma isCont_setsum [simp]: fixes f :: "'a \ 'b::topological_space \ 'c::real_normed_vector" @@ -594,4 +490,29 @@ (X -- a --> (L::'b::topological_space))" using LIMSEQ_SEQ_conv2 LIMSEQ_SEQ_conv1 .. +subsection {* Legacy theorem names *} + +lemmas LIM_ident [simp] = tendsto_ident_at +lemmas LIM_const [simp] = tendsto_const [where F="at x", standard] +lemmas LIM_add = tendsto_add [where F="at x", standard] +lemmas LIM_add_zero = tendsto_add_zero [where F="at x", standard] +lemmas LIM_minus = tendsto_minus [where F="at x", standard] +lemmas LIM_diff = tendsto_diff [where F="at x", standard] +lemmas LIM_norm = tendsto_norm [where F="at x", standard] +lemmas LIM_norm_zero = tendsto_norm_zero [where F="at x", standard] +lemmas LIM_norm_zero_cancel = tendsto_norm_zero_cancel [where F="at x", standard] +lemmas LIM_norm_zero_iff = tendsto_norm_zero_iff [where F="at x", standard] +lemmas LIM_rabs = tendsto_rabs [where F="at x", standard] +lemmas LIM_rabs_zero = tendsto_rabs_zero [where F="at x", standard] +lemmas LIM_rabs_zero_cancel = tendsto_rabs_zero_cancel [where F="at x", standard] +lemmas LIM_rabs_zero_iff = tendsto_rabs_zero_iff [where F="at x", standard] +lemmas LIM_compose = tendsto_compose [where F="at x", standard] +lemmas LIM_mult = tendsto_mult [where F="at x", standard] +lemmas LIM_scaleR = tendsto_scaleR [where F="at x", standard] +lemmas LIM_of_real = tendsto_of_real [where F="at x", standard] +lemmas LIM_power = tendsto_power [where F="at x", standard] +lemmas LIM_inverse = tendsto_inverse [where F="at x", standard] +lemmas LIM_sgn = tendsto_sgn [where F="at x", standard] +lemmas isCont_LIM_compose = isCont_tendsto_compose [where F="at x", standard] + end diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/MacLaurin.thy --- a/src/HOL/MacLaurin.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/MacLaurin.thy Sat Aug 20 09:30:23 2011 +0200 @@ -417,9 +417,6 @@ cos (x + real (m) * pi / 2)" by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto) -lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" - unfolding sin_coeff_def by simp (* TODO: move *) - lemma Maclaurin_sin_expansion2: "\t. abs t \ abs x & sin x = @@ -486,9 +483,6 @@ subsection{*Maclaurin Expansion for Cosine Function*} -lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" - unfolding cos_coeff_def by simp (* TODO: move *) - lemma sumr_cos_zero_one [simp]: "(\m=0..<(Suc n). cos_coeff m * 0 ^ m) = 1" by (induct "n", auto) diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/NSA/HTranscendental.thy --- a/src/HOL/NSA/HTranscendental.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/NSA/HTranscendental.thy Sat Aug 20 09:30:23 2011 +0200 @@ -311,7 +311,7 @@ by transfer (rule exp_ln_iff) lemma starfun_exp_ln_eq: "!!u x. ( *f* exp) u = x ==> ( *f* ln) x = u" -by transfer (rule exp_ln_eq) +by transfer (rule ln_unique) lemma starfun_ln_less_self [simp]: "!!x. 0 < x ==> ( *f* ln) x < x" by transfer (rule ln_less_self) diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/NthRoot.thy --- a/src/HOL/NthRoot.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/NthRoot.thy Sat Aug 20 09:30:23 2011 +0200 @@ -518,12 +518,6 @@ apply (rule real_sqrt_abs) done -lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\" -by simp (* TODO: delete *) - -lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \ 0" -by simp (* TODO: delete *) - lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x" by (simp add: power_inverse [symmetric]) @@ -533,15 +527,6 @@ lemma real_sqrt_ge_one: "1 \ x ==> 1 \ sqrt x" by simp -lemma real_sqrt_two_gt_zero [simp]: "0 < sqrt 2" -by simp - -lemma real_sqrt_two_ge_zero [simp]: "0 \ sqrt 2" -by simp - -lemma real_sqrt_two_gt_one [simp]: "1 < sqrt 2" -by simp - lemma sqrt_divide_self_eq: assumes nneg: "0 \ x" shows "sqrt x / x = inverse (sqrt x)" @@ -575,21 +560,18 @@ subsection {* Square Root of Sum of Squares *} -lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \ sqrt(x*x + y*y)" -by (rule real_sqrt_ge_zero [OF sum_squares_ge_zero]) - -lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \ sqrt (x\ + y\)" -by simp +lemma real_sqrt_sum_squares_ge_zero: "0 \ sqrt (x\ + y\)" + by simp (* TODO: delete *) declare real_sqrt_sum_squares_ge_zero [THEN abs_of_nonneg, simp] lemma real_sqrt_sum_squares_mult_ge_zero [simp]: "0 \ sqrt ((x\ + y\)*(xa\ + ya\))" -by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff) + by (simp add: zero_le_mult_iff) lemma real_sqrt_sum_squares_mult_squared_eq [simp]: "sqrt ((x\ + y\) * (xa\ + ya\)) ^ 2 = (x\ + y\) * (xa\ + ya\)" -by (auto simp add: zero_le_mult_iff) + by (simp add: zero_le_mult_iff) lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\ + y\) = x \ y = 0" by (drule_tac f = "%x. x\" in arg_cong, simp) diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/SEQ.thy Sat Aug 20 09:30:23 2011 +0200 @@ -272,9 +272,6 @@ lemma LIMSEQ_iff_nz: "X ----> L = (\r>0. \no>0. \n\no. dist (X n) L < r)" unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc) -lemma LIMSEQ_Zfun_iff: "((\n. X n) ----> L) = Zfun (\n. X n - L) sequentially" -by (rule tendsto_Zfun_iff) - lemma metric_LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X ----> L" by (simp add: LIMSEQ_def) @@ -293,19 +290,11 @@ shows "\X ----> L; 0 < r\ \ \no. \n\no. norm (X n - L) < r" by (simp add: LIMSEQ_iff) -lemma LIMSEQ_const: "(\n. k) ----> k" -by (rule tendsto_const) - lemma LIMSEQ_const_iff: fixes k l :: "'a::t2_space" shows "(\n. k) ----> l \ k = l" using trivial_limit_sequentially by (rule tendsto_const_iff) -lemma LIMSEQ_norm: - fixes a :: "'a::real_normed_vector" - shows "X ----> a \ (\n. norm (X n)) ----> norm a" -by (rule tendsto_norm) - lemma LIMSEQ_ignore_initial_segment: "f ----> a \ (\n. f (n + k)) ----> a" apply (rule topological_tendstoI) @@ -341,44 +330,11 @@ unfolding tendsto_def eventually_sequentially by (metis div_le_dividend div_mult_self1_is_m le_trans nat_mult_commute) -lemma LIMSEQ_add: - fixes a b :: "'a::real_normed_vector" - shows "\X ----> a; Y ----> b\ \ (\n. X n + Y n) ----> a + b" -by (rule tendsto_add) - -lemma LIMSEQ_minus: - fixes a :: "'a::real_normed_vector" - shows "X ----> a \ (\n. - X n) ----> - a" -by (rule tendsto_minus) - -lemma LIMSEQ_minus_cancel: - fixes a :: "'a::real_normed_vector" - shows "(\n. - X n) ----> - a \ X ----> a" -by (rule tendsto_minus_cancel) - -lemma LIMSEQ_diff: - fixes a b :: "'a::real_normed_vector" - shows "\X ----> a; Y ----> b\ \ (\n. X n - Y n) ----> a - b" -by (rule tendsto_diff) - lemma LIMSEQ_unique: fixes a b :: "'a::t2_space" shows "\X ----> a; X ----> b\ \ a = b" using trivial_limit_sequentially by (rule tendsto_unique) -lemma (in bounded_linear) LIMSEQ: - "X ----> a \ (\n. f (X n)) ----> f a" -by (rule tendsto) - -lemma (in bounded_bilinear) LIMSEQ: - "\X ----> a; Y ----> b\ \ (\n. X n ** Y n) ----> a ** b" -by (rule tendsto) - -lemma LIMSEQ_mult: - fixes a b :: "'a::real_normed_algebra" - shows "[| X ----> a; Y ----> b |] ==> (%n. X n * Y n) ----> a * b" - by (rule tendsto_mult) - lemma increasing_LIMSEQ: fixes f :: "nat \ real" assumes inc: "!!n. f n \ f (Suc n)" @@ -424,33 +380,6 @@ shows "\X ----> a; a \ 0\ \ Bseq (\n. inverse (X n))" unfolding Bseq_conv_Bfun by (rule Bfun_inverse) -lemma LIMSEQ_inverse: - fixes a :: "'a::real_normed_div_algebra" - shows "\X ----> a; a \ 0\ \ (\n. inverse (X n)) ----> inverse a" -by (rule tendsto_inverse) - -lemma LIMSEQ_divide: - fixes a b :: "'a::real_normed_field" - shows "\X ----> a; Y ----> b; b \ 0\ \ (\n. X n / Y n) ----> a / b" -by (rule tendsto_divide) - -lemma LIMSEQ_pow: - fixes a :: "'a::{power, real_normed_algebra}" - shows "X ----> a \ (\n. (X n) ^ m) ----> a ^ m" - by (rule tendsto_power) - -lemma LIMSEQ_setsum: - fixes L :: "'a \ 'b::real_normed_vector" - assumes n: "\n. n \ S \ X n ----> L n" - shows "(\m. \n\S. X n m) ----> (\n\S. L n)" -using assms by (rule tendsto_setsum) - -lemma LIMSEQ_setprod: - fixes L :: "'a \ 'b::{real_normed_algebra,comm_ring_1}" - assumes n: "\n. n \ S \ X n ----> L n" - shows "(\m. \n\S. X n m) ----> (\n\S. L n)" - using assms by (rule tendsto_setprod) - lemma LIMSEQ_add_const: (* FIXME: delete *) fixes a :: "'a::real_normed_vector" shows "f ----> a ==> (%n.(f n + b)) ----> a + b" @@ -470,24 +399,12 @@ lemma LIMSEQ_diff_approach_zero: fixes L :: "'a::real_normed_vector" shows "g ----> L ==> (%x. f x - g x) ----> 0 ==> f ----> L" -by (drule (1) LIMSEQ_add, simp) + by (drule (1) tendsto_add, simp) lemma LIMSEQ_diff_approach_zero2: fixes L :: "'a::real_normed_vector" shows "f ----> L ==> (%x. f x - g x) ----> 0 ==> g ----> L" -by (drule (1) LIMSEQ_diff, simp) - -text{*A sequence tends to zero iff its abs does*} -lemma LIMSEQ_norm_zero: - fixes X :: "nat \ 'a::real_normed_vector" - shows "((\n. norm (X n)) ----> 0) \ (X ----> 0)" - by (rule tendsto_norm_zero_iff) - -lemma LIMSEQ_rabs_zero: "((%n. \f n\) ----> 0) = (f ----> (0::real))" - by (rule tendsto_rabs_zero_iff) - -lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \f n\) ----> \l\" - by (rule tendsto_rabs) + by (drule (1) tendsto_diff, simp) text{*An unbounded sequence's inverse tends to 0*} @@ -517,16 +434,17 @@ lemma LIMSEQ_inverse_real_of_nat_add: "(%n. r + inverse(real(Suc n))) ----> r" -by (cut_tac LIMSEQ_add [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto) + using tendsto_add [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus: "(%n. r + -inverse(real(Suc n))) ----> r" -by (cut_tac LIMSEQ_add_minus [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat], auto) + using LIMSEQ_add_minus [OF tendsto_const LIMSEQ_inverse_real_of_nat] by auto lemma LIMSEQ_inverse_real_of_nat_add_minus_mult: "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r" -by (cut_tac b=1 in - LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_inverse_real_of_nat_add_minus], auto) + using tendsto_mult [OF tendsto_const + LIMSEQ_inverse_real_of_nat_add_minus [of 1]] + by auto lemma LIMSEQ_le_const: "\X ----> (x::real); \N. \n\N. a \ X n\ \ a \ x" @@ -542,7 +460,7 @@ "\X ----> (x::real); \N. \n\N. X n \ a\ \ x \ a" apply (subgoal_tac "- a \ - x", simp) apply (rule LIMSEQ_le_const) -apply (erule LIMSEQ_minus) +apply (erule tendsto_minus) apply simp done @@ -550,7 +468,7 @@ "\X ----> x; Y ----> y; \N. \n\N. X n \ Y n\ \ x \ (y::real)" apply (subgoal_tac "0 \ y - x", simp) apply (rule LIMSEQ_le_const) -apply (erule (1) LIMSEQ_diff) +apply (erule (1) tendsto_diff) apply (simp add: le_diff_eq) done @@ -572,14 +490,14 @@ by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) lemma convergent_const: "convergent (\n. c)" -by (rule convergentI, rule LIMSEQ_const) + by (rule convergentI, rule tendsto_const) lemma convergent_add: fixes X Y :: "nat \ 'a::real_normed_vector" assumes "convergent (\n. X n)" assumes "convergent (\n. Y n)" shows "convergent (\n. X n + Y n)" -using assms unfolding convergent_def by (fast intro: LIMSEQ_add) + using assms unfolding convergent_def by (fast intro: tendsto_add) lemma convergent_setsum: fixes X :: "'a \ nat \ 'b::real_normed_vector" @@ -593,19 +511,19 @@ lemma (in bounded_linear) convergent: assumes "convergent (\n. X n)" shows "convergent (\n. f (X n))" -using assms unfolding convergent_def by (fast intro: LIMSEQ) + using assms unfolding convergent_def by (fast intro: tendsto) lemma (in bounded_bilinear) convergent: assumes "convergent (\n. X n)" and "convergent (\n. Y n)" shows "convergent (\n. X n ** Y n)" -using assms unfolding convergent_def by (fast intro: LIMSEQ) + using assms unfolding convergent_def by (fast intro: tendsto) lemma convergent_minus_iff: fixes X :: "nat \ 'a::real_normed_vector" shows "convergent X \ convergent (\n. - X n)" apply (simp add: convergent_def) -apply (auto dest: LIMSEQ_minus) -apply (drule LIMSEQ_minus, auto) +apply (auto dest: tendsto_minus) +apply (drule tendsto_minus, auto) done lemma lim_le: @@ -661,7 +579,7 @@ case True with top_down and `a ----> x` show ?thesis by auto next case False with `monoseq a`[unfolded monoseq_def] have "\ m. \ n \ m. - a m \ - a n" by auto - hence "- a m \ - x" using top_down[OF LIMSEQ_minus[OF `a ----> x`]] by blast + hence "- a m \ - x" using top_down[OF tendsto_minus[OF `a ----> x`]] by blast hence False using `a m < x` by auto thus ?thesis .. qed @@ -676,7 +594,7 @@ show ?thesis by blast next case False hence "- a m < - x" using `a m \ x` by auto - with when_decided[OF LIMSEQ_minus[OF `a ----> x`] monoseq_minus[OF `monoseq a`], where m2=m] + with when_decided[OF tendsto_minus[OF `a ----> x`] monoseq_minus[OF `monoseq a`], where m2=m] show ?thesis by auto qed qed auto @@ -855,8 +773,8 @@ by (blast intro: const [of 0]) have "X = (\n. X 0)" by (blast intro: ext X) - hence "L = X 0" using LIMSEQ_const [of "X 0"] - by (auto intro: LIMSEQ_unique lim) + hence "L = X 0" using tendsto_const [of "X 0" sequentially] + by (auto intro: LIMSEQ_unique lim) thus ?thesis by (blast intro: eq_refl X) qed @@ -867,7 +785,7 @@ have inc: "incseq (\n. - X n)" using dec by (simp add: decseq_eq_incseq) have "- X n \ - L" - by (blast intro: incseq_le [OF inc] LIMSEQ_minus lim) + by (blast intro: incseq_le [OF inc] tendsto_minus lim) thus ?thesis by simp qed @@ -1187,7 +1105,7 @@ "\0 \ (x::real); x < 1\ \ (\n. x ^ n) ----> 0" proof (cases) assume "x = 0" - hence "(\n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const) + hence "(\n. x ^ Suc n) ----> 0" by (simp add: tendsto_const) thus ?thesis by (rule LIMSEQ_imp_Suc) next assume "0 \ x" and "x \ 0" @@ -1204,14 +1122,14 @@ fixes x :: "'a::{real_normed_algebra_1}" shows "norm x < 1 \ (\n. x ^ n) ----> 0" apply (drule LIMSEQ_realpow_zero [OF norm_ge_zero]) -apply (simp only: LIMSEQ_Zfun_iff, erule Zfun_le) +apply (simp only: tendsto_Zfun_iff, erule Zfun_le) apply (simp add: power_abs norm_power_ineq) done lemma LIMSEQ_divide_realpow_zero: "1 < (x::real) ==> (%n. a / (x ^ n)) ----> 0" -apply (cut_tac a = a and x1 = "inverse x" in - LIMSEQ_mult [OF LIMSEQ_const LIMSEQ_realpow_zero]) +using tendsto_mult [OF tendsto_const [of a] + LIMSEQ_realpow_zero [of "inverse x"]] apply (auto simp add: divide_inverse power_inverse) apply (simp add: inverse_eq_divide pos_divide_less_eq) done @@ -1222,8 +1140,29 @@ by (rule LIMSEQ_realpow_zero [OF abs_ge_zero]) lemma LIMSEQ_rabs_realpow_zero2: "\c\ < (1::real) ==> (%n. c ^ n) ----> 0" -apply (rule LIMSEQ_rabs_zero [THEN iffD1]) +apply (rule tendsto_rabs_zero_cancel) apply (auto intro: LIMSEQ_rabs_realpow_zero simp add: power_abs) done +subsection {* Legacy theorem names *} + +lemmas LIMSEQ_Zfun_iff = tendsto_Zfun_iff [where F=sequentially] +lemmas LIMSEQ_const = tendsto_const [where F=sequentially] +lemmas LIMSEQ_norm = tendsto_norm [where F=sequentially] +lemmas LIMSEQ_add = tendsto_add [where F=sequentially] +lemmas LIMSEQ_minus = tendsto_minus [where F=sequentially] +lemmas LIMSEQ_minus_cancel = tendsto_minus_cancel [where F=sequentially] +lemmas LIMSEQ_diff = tendsto_diff [where F=sequentially] +lemmas (in bounded_linear) LIMSEQ = tendsto [where F=sequentially] +lemmas (in bounded_bilinear) LIMSEQ = tendsto [where F=sequentially] +lemmas LIMSEQ_mult = tendsto_mult [where F=sequentially] +lemmas LIMSEQ_inverse = tendsto_inverse [where F=sequentially] +lemmas LIMSEQ_divide = tendsto_divide [where F=sequentially] +lemmas LIMSEQ_pow = tendsto_power [where F=sequentially] +lemmas LIMSEQ_setsum = tendsto_setsum [where F=sequentially] +lemmas LIMSEQ_setprod = tendsto_setprod [where F=sequentially] +lemmas LIMSEQ_norm_zero = tendsto_norm_zero_iff [where F=sequentially] +lemmas LIMSEQ_rabs_zero = tendsto_rabs_zero_iff [where F=sequentially] +lemmas LIMSEQ_imp_rabs = tendsto_rabs [where F=sequentially] + end diff -r 2b088d74beb3 -r 46d5e7f52ba5 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Sat Aug 20 01:40:22 2011 +0200 +++ b/src/HOL/Transcendental.thy Sat Aug 20 09:30:23 2011 +0200 @@ -871,8 +871,15 @@ apply (simp del: of_real_add) done -lemma isCont_exp [simp]: "isCont exp x" -by (rule DERIV_exp [THEN DERIV_isCont]) +lemma isCont_exp: "isCont exp x" + by (rule DERIV_exp [THEN DERIV_isCont]) + +lemma isCont_exp' [simp]: "isCont f a \ isCont (\x. exp (f x)) a" + by (rule isCont_o2 [OF _ isCont_exp]) + +lemma tendsto_exp [tendsto_intros]: + "(f ---> a) F \ ((\x. exp (f x)) ---> exp a) F" + by (rule isCont_tendsto_compose [OF isCont_exp]) subsubsection {* Properties of the Exponential Function *} @@ -1152,9 +1159,6 @@ lemma ln_less_zero: "\0 < x; x < 1\ \ ln x < 0" by simp -lemma exp_ln_eq: "exp u = x ==> ln x = u" - by (rule ln_unique) (* TODO: delete *) - lemma isCont_ln: "0 < x \ isCont ln x" apply (subgoal_tac "isCont ln (exp (ln x))", simp) apply (rule isCont_inverse_function [where f=exp], simp_all) @@ -1162,7 +1166,7 @@ lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) - apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln]) + apply (erule DERIV_cong [OF DERIV_exp exp_ln]) apply (simp_all add: abs_if isCont_ln) done @@ -1216,6 +1220,20 @@ definition cos :: "real \ real" where "cos = (\x. \n. cos_coeff n * x ^ n)" +lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" + unfolding sin_coeff_def by simp + +lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" + unfolding cos_coeff_def by simp + +lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" + unfolding cos_coeff_def sin_coeff_def + by (simp del: mult_Suc) + +lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" + unfolding cos_coeff_def sin_coeff_def + by (simp del: mult_Suc, auto simp add: odd_Suc_mult_two_ex) + lemma summable_sin: "summable (\n. sin_coeff n * x ^ n)" unfolding sin_coeff_def apply (rule summable_comparison_test [OF _ summable_exp [where x="\x\"]]) @@ -1234,49 +1252,47 @@ lemma cos_converges: "(\n. cos_coeff n * x ^ n) sums cos(x)" unfolding cos_def by (rule summable_cos [THEN summable_sums]) -lemma sin_fdiffs: "diffs sin_coeff = cos_coeff" -unfolding sin_coeff_def cos_coeff_def -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 sin_coeff n = cos_coeff n" -by (simp only: sin_fdiffs) - -lemma cos_fdiffs: "diffs cos_coeff = (\n. - sin_coeff n)" -unfolding sin_coeff_def cos_coeff_def -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 cos_coeff n = - sin_coeff n" -by (simp only: cos_fdiffs) +lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" + by (simp add: diffs_def sin_coeff_Suc real_of_nat_def del: of_nat_Suc) + +lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" + by (simp add: diffs_def cos_coeff_Suc real_of_nat_def del: of_nat_Suc) text{*Now at last we can get the derivatives of exp, sin and cos*} -lemma lemma_sin_minus: "- sin x = (\n. - (sin_coeff n * x ^ n))" -by (auto intro!: sums_unique sums_minus sin_converges) - lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" -unfolding sin_def cos_def -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 + unfolding sin_def cos_def + apply (rule DERIV_cong, rule termdiffs [where K="1 + \x\"]) + apply (simp_all add: diffs_sin_coeff diffs_cos_coeff + summable_minus summable_sin summable_cos) + done lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" -unfolding cos_def -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]) - + unfolding cos_def sin_def + apply (rule DERIV_cong, rule termdiffs [where K="1 + \x\"]) + apply (simp_all add: diffs_sin_coeff diffs_cos_coeff diffs_minus + summable_minus summable_sin summable_cos suminf_minus) + done + +lemma isCont_sin: "isCont sin x" + by (rule DERIV_sin [THEN DERIV_isCont]) + +lemma isCont_cos: "isCont cos x" + by (rule DERIV_cos [THEN DERIV_isCont]) + +lemma isCont_sin' [simp]: "isCont f a \ isCont (\x. sin (f x)) a" + by (rule isCont_o2 [OF _ isCont_sin]) + +lemma isCont_cos' [simp]: "isCont f a \ isCont (\x. cos (f x)) a" + by (rule isCont_o2 [OF _ isCont_cos]) + +lemma tendsto_sin [tendsto_intros]: + "(f ---> a) F \ ((\x. sin (f x)) ---> sin a) F" + by (rule isCont_tendsto_compose [OF isCont_sin]) + +lemma tendsto_cos [tendsto_intros]: + "(f ---> a) F \ ((\x. cos (f x)) ---> cos a) F" + by (rule isCont_tendsto_compose [OF isCont_cos]) declare DERIV_exp[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] @@ -1287,13 +1303,10 @@ subsection {* Properties of Sine and Cosine *} lemma sin_zero [simp]: "sin 0 = 0" -unfolding sin_def sin_coeff_def by (simp add: powser_zero) + unfolding sin_def sin_coeff_def by (simp add: powser_zero) lemma cos_zero [simp]: "cos 0 = 1" -unfolding cos_def cos_coeff_def by (simp add: powser_zero) - -lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" - by (rule DERIV_cong) (* TODO: delete *) + unfolding cos_def cos_coeff_def by (simp add: powser_zero) lemma sin_cos_squared_add [simp]: "(sin x)\ + (cos x)\ = 1" proof - @@ -1336,32 +1349,19 @@ lemma DERIV_fun_pow: "DERIV g x :> m ==> DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" -unfolding One_nat_def -apply (rule DERIV_cong) -apply (rule_tac f = "(%x. x ^ n)" in DERIV_chain2) -apply (rule DERIV_pow, auto) -done + by (auto intro!: DERIV_intros) lemma DERIV_fun_exp: "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m" -apply (rule DERIV_cong) -apply (rule_tac f = exp in DERIV_chain2) -apply (rule DERIV_exp, auto) -done + by (auto intro!: DERIV_intros) lemma DERIV_fun_sin: "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m" -apply (rule DERIV_cong) -apply (rule_tac f = sin in DERIV_chain2) -apply (rule DERIV_sin, auto) -done + by (auto intro!: DERIV_intros) lemma DERIV_fun_cos: "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m" -apply (rule DERIV_cong) -apply (rule_tac f = cos in DERIV_chain2) -apply (rule DERIV_cos, auto) -done + by (auto intro!: DERIV_intros) lemma sin_cos_add_lemma: "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + @@ -1485,10 +1485,10 @@ done lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x" -by (auto intro: sin_gt_zero) + by (rule sin_gt_zero) lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1" -apply (cut_tac x = x in sin_gt_zero1) +apply (cut_tac x = x in sin_gt_zero) apply (auto simp add: cos_squared_eq cos_double) done @@ -1888,59 +1888,41 @@ subsection {* Tangent *} -definition - tan :: "real => real" where - "tan x = (sin x)/(cos x)" +definition tan :: "real \ real" where + "tan = (\x. sin x / cos x)" lemma tan_zero [simp]: "tan 0 = 0" -by (simp add: tan_def) + by (simp add: tan_def) lemma tan_pi [simp]: "tan pi = 0" -by (simp add: tan_def) + by (simp add: tan_def) lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0" -by (simp add: tan_def) + by (simp add: tan_def) lemma tan_minus [simp]: "tan (-x) = - tan x" -by (simp add: tan_def minus_mult_left) + by (simp add: tan_def) lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x" -by (simp add: tan_def) + by (simp add: tan_def) lemma lemma_tan_add1: - "[| cos x \ 0; cos y \ 0 |] - ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)" -apply (simp add: tan_def divide_inverse) -apply (auto simp del: inverse_mult_distrib - simp add: inverse_mult_distrib [symmetric] mult_ac) -apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) -apply (auto simp del: inverse_mult_distrib - simp add: mult_assoc left_diff_distrib cos_add) -done + "\cos x \ 0; cos y \ 0\ \ 1 - tan x * tan y = cos (x + y)/(cos x * cos y)" + by (simp add: tan_def cos_add field_simps) lemma add_tan_eq: - "[| cos x \ 0; cos y \ 0 |] - ==> tan x + tan y = sin(x + y)/(cos x * cos y)" -apply (simp add: tan_def) -apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) -apply (auto simp add: mult_assoc left_distrib) -apply (simp add: sin_add) -done + "\cos x \ 0; cos y \ 0\ \ tan x + tan y = sin(x + y)/(cos x * cos y)" + by (simp add: tan_def sin_add field_simps) lemma tan_add: "[| cos x \ 0; cos y \ 0; cos (x + y) \ 0 |] ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))" -apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1) -apply (simp add: tan_def) -done + by (simp add: add_tan_eq lemma_tan_add1, simp add: tan_def) lemma tan_double: "[| cos x \ 0; cos (2 * x) \ 0 |] ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))" -apply (insert tan_add [of x x]) -apply (simp add: mult_2 [symmetric]) -apply (auto simp add: numeral_2_eq_2) -done + using tan_add [of x x] by (simp add: power2_eq_square) lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x" by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) @@ -1968,26 +1950,23 @@ finally show ?thesis . qed -lemma lemma_DERIV_tan: - "cos x \ 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\)" - by (auto intro!: DERIV_intros simp add: field_simps numeral_2_eq_2) - -lemma DERIV_tan [simp]: "cos x \ 0 ==> DERIV tan x :> inverse((cos x)\)" -by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric]) - -lemma isCont_tan [simp]: "cos x \ 0 ==> isCont tan x" -by (rule DERIV_tan [THEN DERIV_isCont]) - -lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" -apply (subgoal_tac "(\x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1") -apply (simp add: divide_inverse [symmetric]) -apply (rule LIM_mult) -apply (rule_tac [2] inverse_1 [THEN subst]) -apply (rule_tac [2] LIM_inverse) -apply (simp_all add: divide_inverse [symmetric]) -apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) -apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+ -done +lemma DERIV_tan [simp]: "cos x \ 0 \ DERIV tan x :> inverse ((cos x)\)" + unfolding tan_def + by (auto intro!: DERIV_intros, simp add: divide_inverse power2_eq_square) + +lemma isCont_tan: "cos x \ 0 \ isCont tan x" + by (rule DERIV_tan [THEN DERIV_isCont]) + +lemma isCont_tan' [simp]: + "\isCont f a; cos (f a) \ 0\ \ isCont (\x. tan (f x)) a" + by (rule isCont_o2 [OF _ isCont_tan]) + +lemma tendsto_tan [tendsto_intros]: + "\(f ---> a) F; cos a \ 0\ \ ((\x. tan (f x)) ---> tan a) F" + by (rule isCont_tendsto_compose [OF isCont_tan]) + +lemma LIM_cos_div_sin: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" + by (rule LIM_cong_limit, (rule tendsto_intros)+, simp_all) lemma lemma_tan_total: "0 < y ==> \x. 0 < x & x < pi/2 & y < tan x" apply (cut_tac LIM_cos_div_sin) @@ -2421,9 +2400,6 @@ lemma tan_60: "tan (pi / 3) = sqrt 3" unfolding tan_def by (simp add: sin_60 cos_60) -lemma DERIV_sin_add: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)" - by (auto intro!: DERIV_intros) (* TODO: delete *) - lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n" proof - have "sin ((real n + 1/2) * pi) = cos (real n * pi)"