# HG changeset patch # User huffman # Date 1313791630 25200 # Node ID d81d57979771d07954522fa9af2ced648252b4ca # Parent 471ff02a85741c75ea62863f0dc74429db301c9f SEQ.thy: legacy theorem names diff -r 471ff02a8574 -r d81d57979771 src/HOL/SEQ.thy --- a/src/HOL/SEQ.thy Fri Aug 19 14:46:45 2011 -0700 +++ b/src/HOL/SEQ.thy Fri Aug 19 15:07:10 2011 -0700 @@ -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