merged
authorhuffman
Fri, 19 Aug 2011 16:55:43 -0700
changeset 44315 349842366154
parent 44314 dbad46932536 (diff)
parent 44309 d4decbd67703 (current diff)
child 44316 84b6f7a6cea4
merged
--- a/src/HOL/Deriv.thy	Fri Aug 19 23:48:18 2011 +0200
+++ b/src/HOL/Deriv.thy	Fri Aug 19 16:55:43 2011 -0700
@@ -524,7 +524,7 @@
                 ((\<forall>n. l \<le> 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
 
--- a/src/HOL/Lim.thy	Fri Aug 19 23:48:18 2011 +0200
+++ b/src/HOL/Lim.thy	Fri Aug 19 16:55:43 2011 -0700
@@ -81,38 +81,8 @@
   shows "(\<lambda>h. f (a + h)) -- 0 --> L \<Longrightarrow> 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: "\<lbrakk> f -- x --> L ; K = L \<rbrakk> \<Longrightarrow> f -- x --> K" by simp
 
-lemma LIM_add:
-  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  assumes f: "f -- a --> L" and g: "g -- a --> M"
-  shows "(\<lambda>x. f x + g x) -- a --> (L + M)"
-using assms by (rule tendsto_add)
-
-lemma LIM_add_zero:
-  fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "\<lbrakk>f -- a --> 0; g -- a --> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. f x + g x) -- a --> 0"
-  by (rule tendsto_add_zero)
-
-lemma LIM_minus:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "f -- a --> L \<Longrightarrow> (\<lambda>x. - f x) -- a --> - L"
-by (rule tendsto_minus)
-
-(* TODO: delete *)
-lemma LIM_add_minus:
-  fixes f g :: "'a::topological_space \<Rightarrow> '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 \<Rightarrow> 'b::real_normed_vector"
-  shows "\<lbrakk>f -- x --> l; g -- x --> m\<rbrakk> \<Longrightarrow> (\<lambda>x. f x - g x) -- x --> l - m"
-by (rule tendsto_diff)
-
 lemma LIM_zero:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   shows "f -- a --> l \<Longrightarrow> (\<lambda>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 \<Rightarrow> 'b::real_normed_vector"
-  shows "f -- a --> l \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> norm l"
-by (rule tendsto_norm)
-
-lemma LIM_norm_zero:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "f -- a --> 0 \<Longrightarrow> (\<lambda>x. norm (f x)) -- a --> 0"
-by (rule tendsto_norm_zero)
-
-lemma LIM_norm_zero_cancel:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "(\<lambda>x. norm (f x)) -- a --> 0 \<Longrightarrow> f -- a --> 0"
-by (rule tendsto_norm_zero_cancel)
-
-lemma LIM_norm_zero_iff:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
-  shows "(\<lambda>x. norm (f x)) -- a --> 0 = f -- a --> 0"
-by (rule tendsto_norm_zero_iff)
-
-lemma LIM_rabs: "f -- a --> (l::real) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) -- a --> \<bar>l\<bar>"
-  by (rule tendsto_rabs)
-
-lemma LIM_rabs_zero: "f -- a --> (0::real) \<Longrightarrow> (\<lambda>x. \<bar>f x\<bar>) -- a --> 0"
-  by (rule tendsto_rabs_zero)
-
-lemma LIM_rabs_zero_cancel: "(\<lambda>x. \<bar>f x\<bar>) -- a --> (0::real) \<Longrightarrow> f -- a --> 0"
-  by (rule tendsto_rabs_zero_cancel)
-
-lemma LIM_rabs_zero_iff: "(\<lambda>x. \<bar>f x\<bar>) -- a --> (0::real) = f -- a --> 0"
-  by (rule tendsto_rabs_zero_iff)
-
 lemma trivial_limit_at:
   fixes a :: "'a::real_normed_algebra_1"
   shows "\<not> trivial_limit (at a)"  -- {* TODO: find a more appropriate class *}
@@ -203,9 +141,6 @@
   shows "\<lbrakk>f -- a --> L; f -- a --> M\<rbrakk> \<Longrightarrow> L = M"
   using trivial_limit_at by (rule tendsto_unique)
 
-lemma LIM_ident [simp]: "(\<lambda>x. x) -- a --> a"
-by (rule tendsto_ident_at)
-
 text{*Limits are equal for functions equal except at limit point*}
 lemma LIM_equal:
      "[| \<forall>x. x \<noteq> a --> (f x = g x) |] ==> (f -- a --> l) = (g -- a --> l)"
@@ -235,20 +170,6 @@
   shows "g -- a --> l \<Longrightarrow> 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 \<Rightarrow> '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 "(\<lambda>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: "\<exists>d>0. \<forall>x. x \<noteq> a \<and> dist x a < d \<longrightarrow> f x \<noteq> b"
   shows "(\<lambda>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: "\<lbrakk>g -- l --> g l; f -- a --> l\<rbrakk> \<Longrightarrow> (g \<circ> 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 \<Rightarrow> real"
@@ -321,9 +242,6 @@
   "f -- a --> 0 \<Longrightarrow> (\<lambda>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 \<Rightarrow> 'b::{power,real_normed_algebra}"
-  assumes f: "f -- a --> l"
-  shows "(\<lambda>x. f x ^ n) -- a --> l ^ n"
-  using assms by (rule tendsto_power)
-
-lemma LIM_inverse:
-  fixes L :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>f -- a --> L; L \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>x. inverse (f x)) -- a --> inverse L"
-by (rule tendsto_inverse)
-
 lemma LIM_inverse_fun:
   assumes a: "a \<noteq> (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 \<Rightarrow> 'b::real_normed_vector"
-  shows "\<lbrakk>f -- a --> l; l \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>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 (\<lambda>x. x) a"
-  unfolding isCont_def by (rule LIM_ident)
+  unfolding isCont_def by (rule tendsto_ident_at)
 
 lemma isCont_const [simp]: "isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_vector"
   shows "isCont f a \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> real"
   shows "isCont f a \<Longrightarrow> isCont (\<lambda>x. \<bar>f x\<bar>) a"
-  unfolding isCont_def by (rule LIM_rabs)
+  unfolding isCont_def by (rule tendsto_rabs)
 
 lemma isCont_add [simp]:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
   shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_vector"
   shows "isCont f a \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_vector"
   shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_algebra"
   shows "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_div_algebra"
   shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_field"
   shows "\<lbrakk>isCont f a; isCont g a; g a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. f x / g x) a"
   unfolding isCont_def by (rule tendsto_divide)
 
-lemma isCont_LIM_compose:
-  "\<lbrakk>isCont g l; f -- a --> l\<rbrakk> \<Longrightarrow> (\<lambda>x. g (f x)) -- a --> g l"
-  unfolding isCont_def by (rule LIM_compose)
+lemma isCont_tendsto_compose:
+  "\<lbrakk>isCont g l; (f ---> l) F\<rbrakk> \<Longrightarrow> ((\<lambda>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: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. g (f x)) a"
-  unfolding isCont_def by (rule LIM_compose)
+  unfolding isCont_def by (rule tendsto_compose)
 
 lemma isCont_o: "\<lbrakk>isCont f a; isCont g (f a)\<rbrakk> \<Longrightarrow> isCont (g o f) a"
   unfolding o_def by (rule isCont_o2)
 
 lemma (in bounded_linear) isCont:
   "isCont g a \<Longrightarrow> isCont (\<lambda>x. f (g x)) a"
-  unfolding isCont_def by (rule LIM)
+  unfolding isCont_def by (rule tendsto)
 
 lemma (in bounded_bilinear) isCont:
   "\<lbrakk>isCont f a; isCont g a\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::{power,real_normed_algebra}"
   shows "isCont f a \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::real_normed_vector"
   shows "\<lbrakk>isCont f a; f a \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>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 \<Rightarrow> 'b::topological_space \<Rightarrow> '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
--- a/src/HOL/SEQ.thy	Fri Aug 19 23:48:18 2011 +0200
+++ b/src/HOL/SEQ.thy	Fri Aug 19 16:55:43 2011 -0700
@@ -272,9 +272,6 @@
 lemma LIMSEQ_iff_nz: "X ----> L = (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)"
   unfolding LIMSEQ_def by (metis Suc_leD zero_less_Suc)
 
-lemma LIMSEQ_Zfun_iff: "((\<lambda>n. X n) ----> L) = Zfun (\<lambda>n. X n - L) sequentially"
-by (rule tendsto_Zfun_iff)
-
 lemma metric_LIMSEQ_I:
   "(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X ----> L"
 by (simp add: LIMSEQ_def)
@@ -293,19 +290,11 @@
   shows "\<lbrakk>X ----> L; 0 < r\<rbrakk> \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. norm (X n - L) < r"
 by (simp add: LIMSEQ_iff)
 
-lemma LIMSEQ_const: "(\<lambda>n. k) ----> k"
-by (rule tendsto_const)
-
 lemma LIMSEQ_const_iff:
   fixes k l :: "'a::t2_space"
   shows "(\<lambda>n. k) ----> l \<longleftrightarrow> k = l"
   using trivial_limit_sequentially by (rule tendsto_const_iff)
 
-lemma LIMSEQ_norm:
-  fixes a :: "'a::real_normed_vector"
-  shows "X ----> a \<Longrightarrow> (\<lambda>n. norm (X n)) ----> norm a"
-by (rule tendsto_norm)
-
 lemma LIMSEQ_ignore_initial_segment:
   "f ----> a \<Longrightarrow> (\<lambda>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 "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) ----> a + b"
-by (rule tendsto_add)
-
-lemma LIMSEQ_minus:
-  fixes a :: "'a::real_normed_vector"
-  shows "X ----> a \<Longrightarrow> (\<lambda>n. - X n) ----> - a"
-by (rule tendsto_minus)
-
-lemma LIMSEQ_minus_cancel:
-  fixes a :: "'a::real_normed_vector"
-  shows "(\<lambda>n. - X n) ----> - a \<Longrightarrow> X ----> a"
-by (rule tendsto_minus_cancel)
-
-lemma LIMSEQ_diff:
-  fixes a b :: "'a::real_normed_vector"
-  shows "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) ----> a - b"
-by (rule tendsto_diff)
-
 lemma LIMSEQ_unique:
   fixes a b :: "'a::t2_space"
   shows "\<lbrakk>X ----> a; X ----> b\<rbrakk> \<Longrightarrow> a = b"
   using trivial_limit_sequentially by (rule tendsto_unique)
 
-lemma (in bounded_linear) LIMSEQ:
-  "X ----> a \<Longrightarrow> (\<lambda>n. f (X n)) ----> f a"
-by (rule tendsto)
-
-lemma (in bounded_bilinear) LIMSEQ:
-  "\<lbrakk>X ----> a; Y ----> b\<rbrakk> \<Longrightarrow> (\<lambda>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 \<Rightarrow> real"
   assumes inc: "!!n. f n \<le> f (Suc n)"
@@ -424,33 +380,6 @@
   shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> Bseq (\<lambda>n. inverse (X n))"
 unfolding Bseq_conv_Bfun by (rule Bfun_inverse)
 
-lemma LIMSEQ_inverse:
-  fixes a :: "'a::real_normed_div_algebra"
-  shows "\<lbrakk>X ----> a; a \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. inverse (X n)) ----> inverse a"
-by (rule tendsto_inverse)
-
-lemma LIMSEQ_divide:
-  fixes a b :: "'a::real_normed_field"
-  shows "\<lbrakk>X ----> a; Y ----> b; b \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>n. X n / Y n) ----> a / b"
-by (rule tendsto_divide)
-
-lemma LIMSEQ_pow:
-  fixes a :: "'a::{power, real_normed_algebra}"
-  shows "X ----> a \<Longrightarrow> (\<lambda>n. (X n) ^ m) ----> a ^ m"
-  by (rule tendsto_power)
-
-lemma LIMSEQ_setsum:
-  fixes L :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Sum>n\<in>S. X n m) ----> (\<Sum>n\<in>S. L n)"
-using assms by (rule tendsto_setsum)
-
-lemma LIMSEQ_setprod:
-  fixes L :: "'a \<Rightarrow> 'b::{real_normed_algebra,comm_ring_1}"
-  assumes n: "\<And>n. n \<in> S \<Longrightarrow> X n ----> L n"
-  shows "(\<lambda>m. \<Prod>n\<in>S. X n m) ----> (\<Prod>n\<in>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 \<Rightarrow> 'a::real_normed_vector"
-  shows "((\<lambda>n. norm (X n)) ----> 0) \<longleftrightarrow> (X ----> 0)"
-  by (rule tendsto_norm_zero_iff)
-
-lemma LIMSEQ_rabs_zero: "((%n. \<bar>f n\<bar>) ----> 0) = (f ----> (0::real))"
-  by (rule tendsto_rabs_zero_iff)
-
-lemma LIMSEQ_imp_rabs: "f ----> (l::real) ==> (%n. \<bar>f n\<bar>) ----> \<bar>l\<bar>"
-  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:
   "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. a \<le> X n\<rbrakk> \<Longrightarrow> a \<le> x"
@@ -542,7 +460,7 @@
   "\<lbrakk>X ----> (x::real); \<exists>N. \<forall>n\<ge>N. X n \<le> a\<rbrakk> \<Longrightarrow> x \<le> a"
 apply (subgoal_tac "- a \<le> - x", simp)
 apply (rule LIMSEQ_le_const)
-apply (erule LIMSEQ_minus)
+apply (erule tendsto_minus)
 apply simp
 done
 
@@ -550,7 +468,7 @@
   "\<lbrakk>X ----> x; Y ----> y; \<exists>N. \<forall>n\<ge>N. X n \<le> Y n\<rbrakk> \<Longrightarrow> x \<le> (y::real)"
 apply (subgoal_tac "0 \<le> 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 (\<lambda>n. c)"
-by (rule convergentI, rule LIMSEQ_const)
+  by (rule convergentI, rule tendsto_const)
 
 lemma convergent_add:
   fixes X Y :: "nat \<Rightarrow> 'a::real_normed_vector"
   assumes "convergent (\<lambda>n. X n)"
   assumes "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>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 \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
@@ -593,19 +511,19 @@
 lemma (in bounded_linear) convergent:
   assumes "convergent (\<lambda>n. X n)"
   shows "convergent (\<lambda>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 (\<lambda>n. X n)" and "convergent (\<lambda>n. Y n)"
   shows "convergent (\<lambda>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 \<Rightarrow> 'a::real_normed_vector"
   shows "convergent X \<longleftrightarrow> convergent (\<lambda>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 "\<forall> m. \<forall> n \<ge> m. - a m \<le> - a n" by auto
-      hence "- a m \<le> - x" using top_down[OF LIMSEQ_minus[OF `a ----> x`]] by blast
+      hence "- a m \<le> - 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 \<noteq> 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 = (\<lambda>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 (\<lambda>n. - X n)" using dec
     by (simp add: decseq_eq_incseq)
   have "- X n \<le> - 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 @@
   "\<lbrakk>0 \<le> (x::real); x < 1\<rbrakk> \<Longrightarrow> (\<lambda>n. x ^ n) ----> 0"
 proof (cases)
   assume "x = 0"
-  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: LIMSEQ_const)
+  hence "(\<lambda>n. x ^ Suc n) ----> 0" by (simp add: tendsto_const)
   thus ?thesis by (rule LIMSEQ_imp_Suc)
 next
   assume "0 \<le> x" and "x \<noteq> 0"
@@ -1204,14 +1122,14 @@
   fixes x :: "'a::{real_normed_algebra_1}"
   shows "norm x < 1 \<Longrightarrow> (\<lambda>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: "\<bar>c\<bar> < (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
--- a/src/HOL/Transcendental.thy	Fri Aug 19 23:48:18 2011 +0200
+++ b/src/HOL/Transcendental.thy	Fri Aug 19 16:55:43 2011 -0700
@@ -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 \<Longrightarrow> isCont (\<lambda>x. exp (f x)) a"
+  by (rule isCont_o2 [OF _ isCont_exp])
+
+lemma tendsto_exp [tendsto_intros]:
+  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. exp (f x)) ---> exp a) F"
+  by (rule isCont_tendsto_compose [OF isCont_exp])
 
 
 subsubsection {* Properties of the Exponential Function *}
@@ -1271,12 +1278,25 @@
 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])
-
+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 \<Longrightarrow> isCont (\<lambda>x. sin (f x)) a"
+  by (rule isCont_o2 [OF _ isCont_sin])
+
+lemma isCont_cos' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. cos (f x)) a"
+  by (rule isCont_o2 [OF _ isCont_cos])
+
+lemma tendsto_sin [tendsto_intros]:
+  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. sin (f x)) ---> sin a) F"
+  by (rule isCont_tendsto_compose [OF isCont_sin])
+
+lemma tendsto_cos [tendsto_intros]:
+  "(f ---> a) F \<Longrightarrow> ((\<lambda>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,10 +1307,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)
+  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 *)
@@ -1336,32 +1356,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 +1492,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 +1895,41 @@
 
 subsection {* Tangent *}
 
-definition
-  tan :: "real => real" where
-  "tan x = (sin x)/(cos x)"
+definition tan :: "real \<Rightarrow> real" where
+  "tan = (\<lambda>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 \<noteq> 0; cos y \<noteq> 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
+  "\<lbrakk>cos x \<noteq> 0; cos y \<noteq> 0\<rbrakk> \<Longrightarrow> 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 \<noteq> 0; cos y \<noteq> 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
+  "\<lbrakk>cos x \<noteq> 0; cos y \<noteq> 0\<rbrakk> \<Longrightarrow> 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 \<noteq> 0; cos y \<noteq> 0; cos (x + y) \<noteq> 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 \<noteq> 0; cos (2 * x) \<noteq> 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 +1957,23 @@
   finally show ?thesis .
 qed
 
-lemma lemma_DERIV_tan:
-     "cos x \<noteq> 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\<twosuperior>)"
-  by (auto intro!: DERIV_intros simp add: field_simps numeral_2_eq_2)
-
-lemma DERIV_tan [simp]: "cos x \<noteq> 0 ==> DERIV tan x :> inverse((cos x)\<twosuperior>)"
-by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric])
-
-lemma isCont_tan [simp]: "cos x \<noteq> 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 "(\<lambda>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 \<noteq> 0 \<Longrightarrow> DERIV tan x :> inverse ((cos x)\<twosuperior>)"
+  unfolding tan_def
+  by (auto intro!: DERIV_intros, simp add: divide_inverse power2_eq_square)
+
+lemma isCont_tan: "cos x \<noteq> 0 \<Longrightarrow> isCont tan x"
+  by (rule DERIV_tan [THEN DERIV_isCont])
+
+lemma isCont_tan' [simp]:
+  "\<lbrakk>isCont f a; cos (f a) \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. tan (f x)) a"
+  by (rule isCont_o2 [OF _ isCont_tan])
+
+lemma tendsto_tan [tendsto_intros]:
+  "\<lbrakk>(f ---> a) F; cos a \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>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 ==> \<exists>x. 0 < x & x < pi/2 & y < tan x"
 apply (cut_tac LIM_cos_div_sin)