cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
--- a/src/HOL/MacLaurin.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/MacLaurin.thy Tue Mar 18 15:53:48 2014 +0100
@@ -17,10 +17,9 @@
lemma Maclaurin_lemma:
"0 < h ==>
- \<exists>B. f h = (\<Sum>m=0..<n. (j m / real (fact m)) * (h^m)) +
+ \<exists>B. f h = (\<Sum>m<n. (j m / real (fact m)) * (h^m)) +
(B * ((h^n) / real(fact n)))"
-by (rule exI[where x = "(f h - (\<Sum>m=0..<n. (j m / real (fact m)) * h^m)) *
- real(fact n) / (h^n)"]) simp
+by (rule exI[where x = "(f h - (\<Sum>m<n. (j m / real (fact m)) * h^m)) * real(fact n) / (h^n)"]) simp
lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
by arith
@@ -33,20 +32,20 @@
fixes B
assumes DERIV : "\<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
and INIT : "n = Suc k"
- defines "difg \<equiv> (\<lambda>m t. diff m t - ((\<Sum>p = 0..<n - m. diff (m + p) 0 / real (fact p) * t ^ p) +
+ defines "difg \<equiv> (\<lambda>m t. diff m t - ((\<Sum>p<n - m. diff (m + p) 0 / real (fact p) * t ^ p) +
B * (t ^ (n - m) / real (fact (n - m)))))" (is "difg \<equiv> (\<lambda>m t. diff m t - ?difg m t)")
shows "\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (difg m) t :> difg (Suc m) t"
proof (rule allI impI)+
fix m t assume INIT2: "m < n & 0 \<le> t & t \<le> h"
have "DERIV (difg m) t :> diff (Suc m) t -
- ((\<Sum>x = 0..<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) +
+ ((\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) +
real (n - m) * t ^ (n - Suc m) * B / real (fact (n - m)))" unfolding difg_def
by (auto intro!: DERIV_intros DERIV[rule_format, OF INIT2])
- moreover
+ moreover
from INIT2 have intvl: "{..<n - m} = insert 0 (Suc ` {..<n - Suc m})" and "0 < n - m"
unfolding atLeast0LessThan[symmetric] by auto
- have "(\<Sum>x = 0..<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) =
- (\<Sum>x = 0..<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / real (fact (Suc x)))"
+ have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / real (fact x)) =
+ (\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / real (fact (Suc x)))"
unfolding intvl atLeast0LessThan by (subst setsum.insert) (auto simp: setsum.reindex)
moreover
have fact_neq_0: "\<And>x::nat. real (fact x) + real x * real (fact x) \<noteq> 0"
@@ -71,29 +70,26 @@
shows
"\<exists>t. 0 < t & t < h &
f h =
- setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {0..<n} +
+ setsum (%m. (diff m 0 / real (fact m)) * h ^ m) {..<n} +
(diff n t / real (fact n)) * h ^ n"
proof -
from n obtain m where m: "n = Suc m"
by (cases n) (simp add: n)
obtain B where f_h: "f h =
- (\<Sum>m = 0..<n. diff m (0\<Colon>real) / real (fact m) * h ^ m) +
+ (\<Sum>m<n. diff m (0\<Colon>real) / real (fact m) * h ^ m) +
B * (h ^ n / real (fact n))"
using Maclaurin_lemma [OF h] ..
def g \<equiv> "(\<lambda>t. f t -
- (setsum (\<lambda>m. (diff m 0 / real(fact m)) * t^m) {0..<n}
+ (setsum (\<lambda>m. (diff m 0 / real(fact m)) * t^m) {..<n}
+ (B * (t^n / real(fact n)))))"
have g2: "g 0 = 0 & g h = 0"
- apply (simp add: m f_h g_def del: setsum_op_ivl_Suc)
- apply (cut_tac n = m and k = "Suc 0" in sumr_offset2)
- apply (simp add: eq_diff_eq' diff_0 del: setsum_op_ivl_Suc)
- done
+ by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum_reindex)
def difg \<equiv> "(%m t. diff m t -
- (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {0..<n-m}
+ (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {..<n-m}
+ (B * ((t ^ (n - m)) / real (fact (n - m))))))"
have difg_0: "difg 0 = g"
@@ -103,14 +99,8 @@
m < n \<and> (0\<Colon>real) \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2)
- have difg_eq_0: "\<forall>m. m < n --> difg m 0 = 0"
- apply clarify
- apply (simp add: m difg_def)
- apply (frule less_iff_Suc_add [THEN iffD1], clarify)
- apply (simp del: setsum_op_ivl_Suc)
- apply (insert sumr_offset4 [of "Suc 0"])
- apply (simp del: setsum_op_ivl_Suc fact_Suc)
- done
+ have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
+ by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum_reindex)
have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
@@ -166,7 +156,7 @@
show "0 < t" by fact
show "t < h" by fact
show "f h =
- (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) +
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n"
using `difg (Suc m) t = 0`
by (simp add: m f_h difg_def del: fact_Suc)
@@ -177,7 +167,7 @@
"0 < h & n>0 & diff 0 = f &
(\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
--> (\<exists>t. 0 < t & t < h &
- f h = (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
+ f h = (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n)"
by (blast intro: Maclaurin)
@@ -187,7 +177,7 @@
and DERIV: "\<forall>m t.
m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t"
shows "\<exists>t. 0 < t \<and> t \<le> h \<and> f h =
- (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n"
proof (cases "n")
case 0 with INIT1 INIT2 show ?thesis by fastforce
@@ -196,7 +186,7 @@
hence "n > 0" by simp
from INIT1 this INIT2 DERIV have "\<exists>t>0. t < h \<and>
f h =
- (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) + diff n t / real (fact n) * h ^ n"
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) + diff n t / real (fact n) * h ^ n"
by (rule Maclaurin)
thus ?thesis by fastforce
qed
@@ -208,7 +198,7 @@
--> (\<exists>t. 0 < t &
t \<le> h &
f h =
- (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n)"
by (blast intro: Maclaurin2)
@@ -216,7 +206,7 @@
assumes "h < 0" "0 < n" "diff 0 = f"
and DERIV: "\<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t"
shows "\<exists>t. h < t & t < 0 &
- f h = (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
+ f h = (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n"
proof -
txt "Transform @{text ABL'} into @{text DERIV_intros} format."
@@ -224,7 +214,7 @@
from assms
have "\<exists>t>0. t < - h \<and>
f (- (- h)) =
- (\<Sum>m = 0..<n.
+ (\<Sum>m<n.
(- 1) ^ m * diff m (- 0) / real (fact m) * (- h) ^ m) +
(- 1) ^ n * diff n (- t) / real (fact n) * (- h) ^ n"
by (intro Maclaurin) (auto intro!: DERIV_intros DERIV')
@@ -233,12 +223,12 @@
have "-1 ^ n * diff n (- t) * (- h) ^ n / real (fact n) = diff n (- t) * h ^ n / real (fact n)"
by (auto simp add: power_mult_distrib[symmetric])
moreover
- have "(SUM m = 0..<n. -1 ^ m * diff m 0 * (- h) ^ m / real (fact m)) = (SUM m = 0..<n. diff m 0 * h ^ m / real (fact m))"
+ have "(SUM m<n. -1 ^ m * diff m 0 * (- h) ^ m / real (fact m)) = (SUM m<n. diff m 0 * h ^ m / real (fact m))"
by (auto intro: setsum_cong simp add: power_mult_distrib[symmetric])
ultimately have " h < - t \<and>
- t < 0 \<and>
f h =
- (\<Sum>m = 0..<n. diff m 0 / real (fact m) * h ^ m) + diff n (- t) / real (fact n) * h ^ n"
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) + diff n (- t) / real (fact n) * h ^ n"
by auto
thus ?thesis ..
qed
@@ -250,7 +240,7 @@
--> (\<exists>t. h < t &
t < 0 &
f h =
- (\<Sum>m=0..<n. diff m 0 / real (fact m) * h ^ m) +
+ (\<Sum>m<n. diff m 0 / real (fact m) * h ^ m) +
diff n t / real (fact n) * h ^ n)"
by (blast intro: Maclaurin_minus)
@@ -262,7 +252,7 @@
lemma Maclaurin_bi_le_lemma [rule_format]:
"n>0 \<longrightarrow>
diff 0 0 =
- (\<Sum>m = 0..<n. diff m 0 * 0 ^ m / real (fact m)) +
+ (\<Sum>m<n. diff m 0 * 0 ^ m / real (fact m)) +
diff n 0 * 0 ^ n / real (fact n)"
by (induct "n") auto
@@ -271,7 +261,7 @@
and DERIV : "\<forall>m t. m < n & abs t \<le> abs x --> DERIV (diff m) t :> diff (Suc m) t"
shows "\<exists>t. abs t \<le> abs x &
f x =
- (\<Sum>m=0..<n. diff m 0 / real (fact m) * x ^ m) +
+ (\<Sum>m<n. diff m 0 / real (fact m) * x ^ m) +
diff n t / real (fact n) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
proof cases
assume "n = 0" with `diff 0 = f` show ?thesis by force
@@ -303,7 +293,7 @@
assumes INIT1: "diff 0 = f" and INIT2: "0 < n" and INIT3: "x \<noteq> 0"
and DERIV: "\<forall>m x. DERIV (diff m) x :> diff(Suc m) x"
shows "\<exists>t. 0 < abs t & abs t < abs x & f x =
- (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
+ (\<Sum>m<n. (diff m 0 / real (fact m)) * x ^ m) +
(diff n t / real (fact n)) * x ^ n" (is "\<exists>t. _ \<and> _ \<and> f x = ?f x t")
proof (cases rule: linorder_cases)
assume "x = 0" with INIT3 show "?thesis"..
@@ -327,14 +317,14 @@
(\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
x ~= 0 & n > 0
--> (\<exists>t. 0 < abs t & abs t < abs x &
- f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
+ f x = (\<Sum>m<n. (diff m 0 / real (fact m)) * x ^ m) +
(diff n t / real (fact n)) * x ^ n)"
by (blast intro: Maclaurin_all_lt)
lemma Maclaurin_zero [rule_format]:
"x = (0::real)
==> n \<noteq> 0 -->
- (\<Sum>m=0..<n. (diff m (0::real) / real (fact m)) * x ^ m) =
+ (\<Sum>m<n. (diff m (0::real) / real (fact m)) * x ^ m) =
diff 0 0"
by (induct n, auto)
@@ -343,7 +333,7 @@
assumes INIT: "diff 0 = f"
and DERIV: "\<forall>m x. DERIV (diff m) x :> diff (Suc m) x"
shows "\<exists>t. abs t \<le> abs x & f x =
- (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
+ (\<Sum>m<n. (diff m 0 / real (fact m)) * x ^ m) +
(diff n t / real (fact n)) * x ^ n" (is "\<exists>t. _ \<and> f x = ?f x t")
proof cases
assume "n = 0" with INIT show ?thesis by force
@@ -352,7 +342,7 @@
show ?thesis
proof cases
assume "x = 0"
- with `n \<noteq> 0` have "(\<Sum>m = 0..<n. diff m 0 / real (fact m) * x ^ m) = diff 0 0"
+ with `n \<noteq> 0` have "(\<Sum>m<n. diff m 0 / real (fact m) * x ^ m) = diff 0 0"
by (intro Maclaurin_zero) auto
with INIT `x = 0` `n \<noteq> 0` have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0" by force
thus ?thesis ..
@@ -369,7 +359,7 @@
lemma Maclaurin_all_le_objl: "diff 0 = f &
(\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
--> (\<exists>t. abs t \<le> abs x &
- f x = (\<Sum>m=0..<n. (diff m 0 / real (fact m)) * x ^ m) +
+ f x = (\<Sum>m<n. (diff m 0 / real (fact m)) * x ^ m) +
(diff n t / real (fact n)) * x ^ n)"
by (blast intro: Maclaurin_all_le)
@@ -379,14 +369,14 @@
lemma Maclaurin_exp_lt: "[| x ~= 0; n > 0 |]
==> (\<exists>t. 0 < abs t &
abs t < abs x &
- exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
+ exp x = (\<Sum>m<n. (x ^ m) / real (fact m)) +
(exp t / real (fact n)) * x ^ n)"
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto)
lemma Maclaurin_exp_le:
"\<exists>t. abs t \<le> abs x &
- exp x = (\<Sum>m=0..<n. (x ^ m) / real (fact m)) +
+ exp x = (\<Sum>m<n. (x ^ m) / real (fact m)) +
(exp t / real (fact n)) * x ^ n"
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto)
@@ -420,7 +410,7 @@
lemma Maclaurin_sin_expansion2:
"\<exists>t. abs t \<le> abs x &
sin x =
- (\<Sum>m=0..<n. sin_coeff m * x ^ m)
+ (\<Sum>m<n. sin_coeff m * x ^ m)
+ ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = sin and n = n and x = x
and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl)
@@ -440,7 +430,7 @@
lemma Maclaurin_sin_expansion:
"\<exists>t. sin x =
- (\<Sum>m=0..<n. sin_coeff m * x ^ m)
+ (\<Sum>m<n. sin_coeff m * x ^ m)
+ ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (insert Maclaurin_sin_expansion2 [of x n])
apply (blast intro: elim:)
@@ -450,7 +440,7 @@
"[| n > 0; 0 < x |] ==>
\<exists>t. 0 < t & t < x &
sin x =
- (\<Sum>m=0..<n. sin_coeff m * x ^ m)
+ (\<Sum>m<n. sin_coeff m * x ^ m)
+ ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl)
apply safe
@@ -467,7 +457,7 @@
"0 < x ==>
\<exists>t. 0 < t & t \<le> x &
sin x =
- (\<Sum>m=0..<n. sin_coeff m * x ^ m)
+ (\<Sum>m<n. sin_coeff m * x ^ m)
+ ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl)
apply safe
@@ -484,7 +474,7 @@
subsection{*Maclaurin Expansion for Cosine Function*}
lemma sumr_cos_zero_one [simp]:
- "(\<Sum>m=0..<(Suc n). cos_coeff m * 0 ^ m) = 1"
+ "(\<Sum>m<(Suc n). cos_coeff m * 0 ^ m) = 1"
by (induct "n", auto)
lemma cos_expansion_lemma:
@@ -494,14 +484,14 @@
lemma Maclaurin_cos_expansion:
"\<exists>t. abs t \<le> abs x &
cos x =
- (\<Sum>m=0..<n. cos_coeff m * x ^ m)
+ (\<Sum>m<n. cos_coeff m * x ^ m)
+ ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
apply safe
apply (simp (no_asm))
apply (simp (no_asm) add: cos_expansion_lemma)
apply (case_tac "n", simp)
-apply (simp del: setsum_op_ivl_Suc)
+apply (simp del: setsum_lessThan_Suc)
apply (rule ccontr, simp)
apply (drule_tac x = x in spec, simp)
apply (erule ssubst)
@@ -514,7 +504,7 @@
"[| 0 < x; n > 0 |] ==>
\<exists>t. 0 < t & t < x &
cos x =
- (\<Sum>m=0..<n. cos_coeff m * x ^ m)
+ (\<Sum>m<n. cos_coeff m * x ^ m)
+ ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl)
apply safe
@@ -530,7 +520,7 @@
"[| x < 0; n > 0 |] ==>
\<exists>t. x < t & t < 0 &
cos x =
- (\<Sum>m=0..<n. cos_coeff m * x ^ m)
+ (\<Sum>m<n. cos_coeff m * x ^ m)
+ ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl)
apply safe
@@ -551,7 +541,7 @@
by auto
lemma Maclaurin_sin_bound:
- "abs(sin x - (\<Sum>m=0..<n. sin_coeff m * x ^ m))
+ "abs(sin x - (\<Sum>m<n. sin_coeff m * x ^ m))
\<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
proof -
have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
@@ -567,7 +557,7 @@
done
from Maclaurin_all_le [OF diff_0 DERIV_diff]
obtain t where t1: "\<bar>t\<bar> \<le> \<bar>x\<bar>" and
- t2: "sin x = (\<Sum>m = 0..<n. ?diff m 0 / real (fact m) * x ^ m) +
+ t2: "sin x = (\<Sum>m<n. ?diff m 0 / real (fact m) * x ^ m) +
?diff n t / real (fact n) * x ^ n" by fast
have diff_m_0:
"\<And>m. ?diff m 0 = (if even m then 0
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Mar 18 15:53:48 2014 +0100
@@ -23,7 +23,7 @@
lemma setsum_mult_product:
"setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
- unfolding sumr_group[of h B A, unfolded atLeast0LessThan, symmetric]
+ unfolding setsum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
proof (rule setsum_cong, simp, rule setsum_reindex_cong)
fix i
show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Tue Mar 18 15:53:48 2014 +0100
@@ -1990,7 +1990,7 @@
fixes f :: "nat \<Rightarrow> 'm::euclidean_space \<Rightarrow> 'n::euclidean_space"
assumes "convex s"
and "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_derivative (f' n x)) (at x within s)"
- and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (setsum (\<lambda>i. f' i x h) {0..<n} - g' x h) \<le> e * norm h"
+ and "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. norm (setsum (\<lambda>i. f' i x h) {..<n} - g' x h) \<le> e * norm h"
and "x \<in> s"
and "(\<lambda>n. f n x) sums l"
shows "\<exists>g. \<forall>x\<in>s. (\<lambda>n. f n x) sums (g x) \<and> (g has_derivative g' x) (at x within s)"
--- a/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Mar 18 15:53:48 2014 +0100
@@ -6006,14 +6006,14 @@
done
qed
have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
- norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {0..N+1}"
+ norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}"
unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
apply (rule order_trans)
apply (rule norm_setsum)
apply (subst sum_sum_product)
prefer 3
proof (rule **, safe)
- show "finite {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i}"
+ show "finite {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i}"
apply (rule finite_product_dependent)
using q
apply auto
@@ -6068,7 +6068,7 @@
using nfx True
by (auto simp add: field_simps)
qed
- ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {0..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le>
+ ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le>
(real y + 1) * (content k *\<^sub>R indicator s x)"
apply (rule_tac x=n in exI)
apply safe
@@ -6078,7 +6078,7 @@
apply auto
done
qed (insert as, auto)
- also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {0..N+1}"
+ also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}"
apply (rule setsum_mono)
proof -
case goal1
@@ -6092,8 +6092,8 @@
also have "\<dots> < e * inverse 2 * 2"
unfolding divide_inverse setsum_right_distrib[symmetric]
apply (rule mult_strict_left_mono)
- unfolding power_inverse atLeastLessThanSuc_atLeastAtMost[symmetric]
- apply (subst sumr_geometric)
+ unfolding power_inverse lessThan_Suc_atMost[symmetric]
+ apply (subst geometric_sum)
using goal1
apply auto
done
--- a/src/HOL/Probability/Lebesgue_Integration.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Tue Mar 18 15:53:48 2014 +0100
@@ -2231,21 +2231,21 @@
obtain x where abs_sum: "(\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M)) sums x"
using sums unfolding summable_def ..
- have 1: "\<And>n. integrable M (\<lambda>x. \<Sum>i = 0..<n. f i x)"
+ have 1: "\<And>n. integrable M (\<lambda>x. \<Sum>i<n. f i x)"
using integrable by auto
- have 2: "\<And>j. AE x in M. \<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x"
+ have 2: "\<And>j. AE x in M. \<bar>\<Sum>i<j. f i x\<bar> \<le> ?w x"
using AE_space
proof eventually_elim
fix j x assume [simp]: "x \<in> space M"
- have "\<bar>\<Sum>i = 0..< j. f i x\<bar> \<le> (\<Sum>i = 0..< j. \<bar>f i x\<bar>)" by (rule setsum_abs)
+ have "\<bar>\<Sum>i<j. f i x\<bar> \<le> (\<Sum>i<j. \<bar>f i x\<bar>)" by (rule setsum_abs)
also have "\<dots> \<le> w x" using w[of x] series_pos_le[of "\<lambda>i. \<bar>f i x\<bar>"] unfolding sums_iff by auto
- finally show "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp
+ finally show "\<bar>\<Sum>i<j. f i x\<bar> \<le> ?w x" by simp
qed
have 3: "integrable M ?w"
proof (rule integral_monotone_convergence(1))
- let ?F = "\<lambda>n y. (\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
+ let ?F = "\<lambda>n y. (\<Sum>i<n. \<bar>f i y\<bar>)"
let ?w' = "\<lambda>n y. if y \<in> space M then ?F n y else 0"
have "\<And>n. integrable M (?F n)"
using integrable by (auto intro!: integrable_abs)
@@ -2254,15 +2254,15 @@
by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
show "AE x in M. (\<lambda>n. ?w' n x) ----> ?w x"
using w by (simp_all add: tendsto_const sums_def)
- have *: "\<And>n. integral\<^sup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
+ have *: "\<And>n. integral\<^sup>L M (?w' n) = (\<Sum>i< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
using integrable by (simp add: integrable_abs cong: integral_cong)
from abs_sum
show "(\<lambda>i. integral\<^sup>L M (?w' i)) ----> x" unfolding * sums_def .
qed (simp add: w_borel measurable_If_set)
from summable[THEN summable_rabs_cancel]
- have 4: "AE x in M. (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
- by (auto intro: summable_sumr_LIMSEQ_suminf)
+ have 4: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) ----> (\<Sum>i. f i x)"
+ by (auto intro: summable_LIMSEQ)
note int = integral_dominated_convergence(1,3)[OF 1 2 3 4
borel_measurable_suminf[OF integrableD(1)[OF integrable]]]
--- a/src/HOL/Probability/Lebesgue_Measure.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Mar 18 15:53:48 2014 +0100
@@ -194,7 +194,7 @@
ultimately show ?case
using Suc A by (simp add: Integration.integral_add[symmetric])
qed auto }
- ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"
+ ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV"
by (simp add: atLeast0LessThan)
qed
qed
--- a/src/HOL/Probability/Measure_Space.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Probability/Measure_Space.thy Tue Mar 18 15:53:48 2014 +0100
@@ -16,7 +16,7 @@
"f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x"
unfolding sums_def
apply (subst LIMSEQ_Suc_iff[symmetric])
- unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost ..
+ unfolding lessThan_Suc_atMost ..
subsection "Relate extended reals and the indicator function"
@@ -304,12 +304,12 @@
with count_sum[THEN spec, of "disjointed A"] A(3)
have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
- moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
+ moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
using f(1)[unfolded positive_def] dA
- by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
+ by (auto intro!: summable_LIMSEQ summable_ereal_pos)
from LIMSEQ_Suc[OF this]
have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
- unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
+ unfolding lessThan_Suc_atMost .
moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
using disjointed_additive[OF f A(1,2)] .
ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp
--- a/src/HOL/Probability/Projective_Limit.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Probability/Projective_Limit.thy Tue Mar 18 15:53:48 2014 +0100
@@ -355,10 +355,10 @@
have "(\<Sum>i\<in>{1..n}. 2 powr - real i) = (\<Sum>i\<in>{1..<Suc n}. (1/2) ^ i)"
by (rule setsum_cong)
(auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide)
- also have "{1..<Suc n} = {0..<Suc n} - {0}" by auto
- also have "setsum (op ^ (1 / 2::real)) ({0..<Suc n} - {0}) =
- setsum (op ^ (1 / 2)) ({0..<Suc n}) - 1" by (auto simp: setsum_diff1)
- also have "\<dots> < 1" by (subst sumr_geometric) auto
+ also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
+ also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
+ setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
+ also have "\<dots> < 1" by (subst geometric_sum) auto
finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" .
qed (auto simp:
`0 < ?a` `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` ereal_less_real_iff zero_ereal_def[symmetric])
--- a/src/HOL/Probability/Regularity.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Probability/Regularity.thy Tue Mar 18 15:53:48 2014 +0100
@@ -351,9 +351,9 @@
case (union D)
then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed)
with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure)
- also have "(\<lambda>n. \<Sum>i\<in>{0..<n}. M (D i)) ----> (\<Sum>i. M (D i))"
- by (intro summable_sumr_LIMSEQ_suminf summable_ereal_pos emeasure_nonneg)
- finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i = 0..<n. measure M (D i)) ----> measure M (\<Union>i. D i)"
+ also have "(\<lambda>n. \<Sum>i<n. M (D i)) ----> (\<Sum>i. M (D i))"
+ by (intro summable_LIMSEQ summable_ereal_pos emeasure_nonneg)
+ finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) ----> measure M (\<Union>i. D i)"
by (simp add: emeasure_eq_measure)
have "(\<Union>i. D i) \<in> sets M" using `range D \<subseteq> sets M` by auto
@@ -362,18 +362,17 @@
proof (rule approx_inner)
fix e::real assume "e > 0"
with measure_LIMSEQ
- have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i = 0..<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2"
+ have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2"
by (auto simp: LIMSEQ_def dist_real_def simp del: less_divide_eq_numeral1)
- hence "\<exists>n0. \<bar>(\<Sum>i = 0..<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
- then obtain n0 where n0: "\<bar>(\<Sum>i = 0..<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
+ hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
+ then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
unfolding choice_iff by blast
- have "ereal (\<Sum>i = 0..<n0. measure M (D i)) = (\<Sum>i = 0..<n0. M (D i))"
+ have "ereal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
by (auto simp add: emeasure_eq_measure)
- also have "\<dots> = (\<Sum>i<n0. M (D i))" by (rule setsum_cong) auto
also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule suminf_upper) (auto simp: emeasure_nonneg)
also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
- finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i = 0..<n0. measure M (D i)) < e/2"
+ finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
using n0 by auto
have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
proof
@@ -388,20 +387,20 @@
then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)"
"\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
unfolding choice_iff by blast
- let ?K = "\<Union>i\<in>{0..<n0}. K i"
- have "disjoint_family_on K {0..<n0}" using K `disjoint_family D`
+ let ?K = "\<Union>i\<in>{..<n0}. K i"
+ have "disjoint_family_on K {..<n0}" using K `disjoint_family D`
unfolding disjoint_family_on_def by blast
- hence mK: "measure M ?K = (\<Sum>i = 0..<n0. measure M (K i))" using K
+ hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K
by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
- have "measure M (\<Union>i. D i) < (\<Sum>i = 0..<n0. measure M (D i)) + e/2" using n0 by simp
- also have "(\<Sum>i = 0..<n0. measure M (D i)) \<le> (\<Sum>i = 0..<n0. measure M (K i) + e/(2*Suc n0))"
+ have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
+ also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
using K by (auto intro: setsum_mono simp: emeasure_eq_measure)
- also have "\<dots> = (\<Sum>i = 0..<n0. measure M (K i)) + (\<Sum>i = 0..<n0. e/(2*Suc n0))"
+ also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
by (simp add: setsum.distrib)
- also have "\<dots> \<le> (\<Sum>i = 0..<n0. measure M (K i)) + e / 2" using `0 < e`
+ also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using `0 < e`
by (auto simp: real_of_nat_def[symmetric] field_simps intro!: mult_left_mono)
finally
- have "measure M (\<Union>i. D i) < (\<Sum>i = 0..<n0. measure M (K i)) + e / 2 + e / 2"
+ have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
by auto
hence "M (\<Union>i. D i) < M ?K + e" by (auto simp: mK emeasure_eq_measure)
moreover
--- a/src/HOL/Series.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Series.thy Tue Mar 18 15:53:48 2014 +0100
@@ -7,122 +7,109 @@
Additional contributions by Jeremy Avigad
*)
-header{*Finite Summation and Infinite Series*}
+header {* Finite Summation and Infinite Series *}
theory Series
imports Limits
begin
-definition
- sums :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a \<Rightarrow> bool"
- (infixr "sums" 80) where
- "f sums s = (%n. setsum f {0..<n}) ----> s"
-
-definition
- summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> bool" where
- "summable f = (\<exists>s. f sums s)"
-
-definition
- suminf :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a" where
- "suminf f = (THE s. f sums s)"
-
-notation suminf (binder "\<Sum>" 10)
-
-
-lemma [trans]: "f=g ==> g sums z ==> f sums z"
- by simp
+(* TODO: MOVE *)
+lemma Suc_less_iff: "Suc n < m \<longleftrightarrow> (\<exists>m'. m = Suc m' \<and> n < m')"
+ by (cases m) auto
-lemma sumr_diff_mult_const:
- "setsum f {0..<n} - (real n*r) = setsum (%i. f i - r) {0..<n::nat}"
- by (simp add: setsum_subtractf real_of_nat_def)
-
-lemma real_setsum_nat_ivl_bounded:
- "(!!p. p < n \<Longrightarrow> f(p) \<le> K)
- \<Longrightarrow> setsum f {0..<n::nat} \<le> real n * K"
-using setsum_bounded[where A = "{0..<n}"]
-by (auto simp:real_of_nat_def)
-
-(* Generalize from real to some algebraic structure? *)
-lemma sumr_minus_one_realpow_zero [simp]:
- "(\<Sum>i=0..<2*n. (-1) ^ Suc i) = (0::real)"
-by (induct "n", auto)
-
-(* FIXME this is an awful lemma! *)
-lemma sumr_one_lb_realpow_zero [simp]:
- "(\<Sum>n=Suc 0..<n. f(n) * (0::real) ^ n) = 0"
-by (rule setsum_0', simp)
-
-lemma sumr_group:
- "(\<Sum>m=0..<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {0 ..< n * k}"
-apply (subgoal_tac "k = 0 | 0 < k", auto)
-apply (induct "n")
-apply (simp_all add: setsum_add_nat_ivl add_commute)
+(* TODO: MOVE *)
+lemma norm_ratiotest_lemma:
+ fixes x y :: "'a::real_normed_vector"
+ shows "\<lbrakk>c \<le> 0; norm x \<le> c * norm y\<rbrakk> \<Longrightarrow> x = 0"
+apply (subgoal_tac "norm x \<le> 0", simp)
+apply (erule order_trans)
+apply (simp add: mult_le_0_iff)
done
-lemma sumr_offset3:
- "setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)) + setsum f {0..<k}"
-apply (subst setsum_shift_bounds_nat_ivl [symmetric])
-apply (simp add: setsum_add_nat_ivl add_commute)
+(* TODO: MOVE *)
+lemma rabs_ratiotest_lemma: "[| c \<le> 0; abs x \<le> c * abs y |] ==> x = (0::real)"
+by (erule norm_ratiotest_lemma, simp)
+
+(* TODO: MOVE *)
+lemma le_Suc_ex: "(k::nat) \<le> l ==> (\<exists>n. l = k + n)"
+apply (drule le_imp_less_or_eq)
+apply (auto dest: less_imp_Suc_add)
done
-lemma sumr_offset:
- fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
- shows "(\<Sum>m=0..<n. f(m+k)) = setsum f {0..<n+k} - setsum f {0..<k}"
-by (simp add: sumr_offset3)
+(* MOVE *)
+lemma setsum_even_minus_one [simp]: "(\<Sum>i<2 * n. (-1) ^ Suc i) = (0::'a::ring_1)"
+ by (induct "n") auto
+
+(* MOVE *)
+lemma setsum_nat_group: "(\<Sum>m<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {..< n * k}"
+ apply (subgoal_tac "k = 0 | 0 < k", auto)
+ apply (induct "n")
+ apply (simp_all add: setsum_add_nat_ivl add_commute atLeast0LessThan[symmetric])
+ done
-lemma sumr_offset2:
- "\<forall>f. (\<Sum>m=0..<n::nat. f(m+k)::real) = setsum f {0..<n+k} - setsum f {0..<k}"
-by (simp add: sumr_offset)
+(* MOVE *)
+lemma norm_setsum:
+ fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+ shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"
+ apply (case_tac "finite A")
+ apply (erule finite_induct)
+ apply simp
+ apply simp
+ apply (erule order_trans [OF norm_triangle_ineq add_left_mono])
+ apply simp
+ done
-lemma sumr_offset4:
- "\<forall>n f. setsum f {0::nat..<n+k} = (\<Sum>m=0..<n. f (m+k)::real) + setsum f {0..<k}"
-by (clarify, rule sumr_offset3)
+(* MOVE *)
+lemma norm_bound_subset:
+ fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+ assumes "finite s" "t \<subseteq> s"
+ assumes le: "(\<And>x. x \<in> s \<Longrightarrow> norm(f x) \<le> g x)"
+ shows "norm (setsum f t) \<le> setsum g s"
+proof -
+ have "norm (setsum f t) \<le> (\<Sum>i\<in>t. norm (f i))"
+ by (rule norm_setsum)
+ also have "\<dots> \<le> (\<Sum>i\<in>t. g i)"
+ using assms by (auto intro!: setsum_mono)
+ also have "\<dots> \<le> setsum g s"
+ using assms order.trans[OF norm_ge_zero le]
+ by (auto intro!: setsum_mono3)
+ finally show ?thesis .
+qed
-subsection{* Infinite Sums, by the Properties of Limits*}
+(* MOVE *)
+lemma (in linorder) lessThan_minus_lessThan [simp]:
+ "{..< n} - {..< m} = {m ..< n}"
+ by auto
+
+definition
+ sums :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a \<Rightarrow> bool"
+ (infixr "sums" 80)
+where
+ "f sums s \<longleftrightarrow> (\<lambda>n. \<Sum>i<n. f i) ----> s"
-(*----------------------
- suminf is the sum
- ---------------------*)
-lemma sums_summable: "f sums l ==> summable f"
+definition summable :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> bool" where
+ "summable f \<longleftrightarrow> (\<exists>s. f sums s)"
+
+definition
+ suminf :: "(nat \<Rightarrow> 'a::{topological_space, comm_monoid_add}) \<Rightarrow> 'a"
+ (binder "\<Sum>" 10)
+where
+ "suminf f = (THE s. f sums s)"
+
+lemma sums_subst[trans]: "f = g \<Longrightarrow> g sums z \<Longrightarrow> f sums z"
+ by simp
+
+lemma sums_summable: "f sums l \<Longrightarrow> summable f"
by (simp add: sums_def summable_def, blast)
-lemma summable_sums:
- fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
- assumes "summable f"
- shows "f sums (suminf f)"
-proof -
- from assms obtain s where s: "(\<lambda>n. setsum f {0..<n}) ----> s"
- unfolding summable_def sums_def [abs_def] ..
- then show ?thesis unfolding sums_def [abs_def] suminf_def
- by (rule theI, auto intro!: tendsto_unique[OF trivial_limit_sequentially])
-qed
+lemma summable_iff_convergent: "summable f \<longleftrightarrow> convergent (\<lambda>n. \<Sum>i<n. f i)"
+ by (simp add: summable_def sums_def convergent_def)
-lemma summable_sumr_LIMSEQ_suminf:
- fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
- shows "summable f \<Longrightarrow> (\<lambda>n. setsum f {0..<n}) ----> suminf f"
-by (rule summable_sums [unfolded sums_def])
-
-lemma suminf_eq_lim: "suminf f = lim (%n. setsum f {0..<n})"
+lemma suminf_eq_lim: "suminf f = lim (\<lambda>n. \<Sum>i<n. f i)"
by (simp add: suminf_def sums_def lim_def)
-(*-------------------
- sum is unique
- ------------------*)
-lemma sums_unique:
- fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
- shows "f sums s \<Longrightarrow> (s = suminf f)"
-apply (frule sums_summable[THEN summable_sums])
-apply (auto intro!: tendsto_unique[OF trivial_limit_sequentially] simp add: sums_def)
-done
-
-lemma sums_iff:
- fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
- shows "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"
- by (metis summable_sums sums_summable sums_unique)
-
lemma sums_finite:
- assumes [simp]: "finite N"
- assumes f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
+ assumes [simp]: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
shows "f sums (\<Sum>n\<in>N. f n)"
proof -
{ fix n
@@ -146,266 +133,76 @@
(simp add: eq atLeast0LessThan tendsto_const del: add_Suc_right)
qed
-lemma suminf_finite:
- fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add,t2_space}"
- assumes N: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
- shows "suminf f = (\<Sum>n\<in>N. f n)"
- using sums_finite[OF assms, THEN sums_unique] by simp
-
-lemma sums_If_finite_set:
- "finite A \<Longrightarrow> (\<lambda>r. if r \<in> A then f r else 0 :: 'a::{comm_monoid_add,t2_space}) sums (\<Sum>r\<in>A. f r)"
+lemma sums_If_finite_set: "finite A \<Longrightarrow> (\<lambda>r. if r \<in> A then f r else 0) sums (\<Sum>r\<in>A. f r)"
using sums_finite[of A "(\<lambda>r. if r \<in> A then f r else 0)"] by simp
-lemma sums_If_finite:
- "finite {r. P r} \<Longrightarrow> (\<lambda>r. if P r then f r else 0 :: 'a::{comm_monoid_add,t2_space}) sums (\<Sum>r | P r. f r)"
- using sums_If_finite_set[of "{r. P r}" f] by simp
-
-lemma sums_single:
- "(\<lambda>r. if r = i then f r else 0::'a::{comm_monoid_add,t2_space}) sums f i"
- using sums_If_finite[of "\<lambda>r. r = i" f] by simp
-
-lemma sums_split_initial_segment:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- shows "f sums s ==> (\<lambda>n. f(n + k)) sums (s - (SUM i = 0..< k. f i))"
- apply (unfold sums_def)
- apply (simp add: sumr_offset)
- apply (rule tendsto_diff [OF _ tendsto_const])
- apply (rule LIMSEQ_ignore_initial_segment)
- apply assumption
-done
-
-lemma summable_ignore_initial_segment:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- shows "summable f ==> summable (%n. f(n + k))"
- apply (unfold summable_def)
- apply (auto intro: sums_split_initial_segment)
-done
-
-lemma suminf_minus_initial_segment:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- shows "summable f ==>
- suminf f = s ==> suminf (%n. f(n + k)) = s - (SUM i = 0..< k. f i)"
- apply (frule summable_ignore_initial_segment)
- apply (rule sums_unique [THEN sym])
- apply (frule summable_sums)
- apply (rule sums_split_initial_segment)
- apply auto
-done
+lemma sums_If_finite: "finite {r. P r} \<Longrightarrow> (\<lambda>r. if P r then f r else 0) sums (\<Sum>r | P r. f r)"
+ using sums_If_finite_set[of "{r. P r}"] by simp
-lemma suminf_split_initial_segment:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- shows "summable f ==>
- suminf f = (SUM i = 0..< k. f i) + (\<Sum>n. f(n + k))"
-by (auto simp add: suminf_minus_initial_segment)
-
-lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable a"
- shows "\<exists> N. \<forall> n \<ge> N. \<bar> \<Sum> i. a (i + n) \<bar> < r"
-proof -
- from LIMSEQ_D[OF summable_sumr_LIMSEQ_suminf[OF `summable a`] `0 < r`]
- obtain N :: nat where "\<forall> n \<ge> N. norm (setsum a {0..<n} - suminf a) < r" by auto
- thus ?thesis unfolding suminf_minus_initial_segment[OF `summable a` refl] abs_minus_commute real_norm_def
- by auto
-qed
+lemma sums_single: "(\<lambda>r. if r = i then f r else 0) sums f i"
+ using sums_If_finite[of "\<lambda>r. r = i"] by simp
-lemma sums_Suc:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- assumes sumSuc: "(\<lambda> n. f (Suc n)) sums l" shows "f sums (l + f 0)"
-proof -
- from sumSuc[unfolded sums_def]
- have "(\<lambda>i. \<Sum>n = Suc 0..<Suc i. f n) ----> l" unfolding setsum_reindex[OF inj_Suc] image_Suc_atLeastLessThan[symmetric] comp_def .
- from tendsto_add[OF this tendsto_const, where b="f 0"]
- have "(\<lambda>i. \<Sum>n = 0..<Suc i. f n) ----> l + f 0" unfolding add_commute setsum_head_upt_Suc[OF zero_less_Suc] .
- thus ?thesis unfolding sums_def by (rule LIMSEQ_imp_Suc)
-qed
-
-lemma series_zero:
- fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
- assumes "\<forall>m. n \<le> m \<longrightarrow> f m = 0"
- shows "f sums (setsum f {0..<n})"
-proof -
- { fix k :: nat have "setsum f {0..<k + n} = setsum f {0..<n}"
- using assms by (induct k) auto }
- note setsum_const = this
- show ?thesis
- unfolding sums_def
- apply (rule LIMSEQ_offset[of _ n])
- unfolding setsum_const
- apply (rule tendsto_const)
- done
-qed
+lemma series_zero: (* REMOVE *)
+ "(\<And>m. n \<le> m \<Longrightarrow> f m = 0) \<Longrightarrow> f sums (\<Sum>i<n. f i)"
+ by (rule sums_finite) auto
lemma sums_zero[simp, intro]: "(\<lambda>n. 0) sums 0"
unfolding sums_def by (simp add: tendsto_const)
lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"
-by (rule sums_zero [THEN sums_summable])
+ by (rule sums_zero [THEN sums_summable])
+
+lemma sums_group: "f sums s \<Longrightarrow> 0 < k \<Longrightarrow> (\<lambda>n. setsum f {n * k ..< n * k + k}) sums s"
+ apply (simp only: sums_def setsum_nat_group tendsto_def eventually_sequentially)
+ apply safe
+ apply (erule_tac x=S in allE)
+ apply safe
+ apply (rule_tac x="N" in exI, safe)
+ apply (drule_tac x="n*k" in spec)
+ apply (erule mp)
+ apply (erule order_trans)
+ apply simp
+ done
+
+context
+ fixes f :: "nat \<Rightarrow> 'a::{t2_space, comm_monoid_add}"
+begin
+
+lemma summable_sums[intro]: "summable f \<Longrightarrow> f sums (suminf f)"
+ by (simp add: summable_def sums_def suminf_def)
+ (metis convergent_LIMSEQ_iff convergent_def lim_def)
+
+lemma summable_LIMSEQ: "summable f \<Longrightarrow> (\<lambda>n. \<Sum>i<n. f i) ----> suminf f"
+ by (rule summable_sums [unfolded sums_def])
+
+lemma sums_unique: "f sums s \<Longrightarrow> s = suminf f"
+ by (metis limI suminf_eq_lim sums_def)
+
+lemma sums_iff: "f sums x \<longleftrightarrow> summable f \<and> (suminf f = x)"
+ by (metis summable_sums sums_summable sums_unique)
+
+lemma suminf_finite:
+ assumes N: "finite N" and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
+ shows "suminf f = (\<Sum>n\<in>N. f n)"
+ using sums_finite[OF assms, THEN sums_unique] by simp
+
+end
lemma suminf_zero[simp]: "suminf (\<lambda>n. 0::'a::{t2_space, comm_monoid_add}) = 0"
-by (rule sums_zero [THEN sums_unique, symmetric])
-
-lemma (in bounded_linear) sums:
- "(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"
- unfolding sums_def by (drule tendsto, simp only: setsum)
-
-lemma (in bounded_linear) summable:
- "summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"
-unfolding summable_def by (auto intro: sums)
-
-lemma (in bounded_linear) suminf:
- "summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"
-by (intro sums_unique sums summable_sums)
-
-lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real]
-lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real]
-lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real]
-
-lemma sums_mult:
- fixes c :: "'a::real_normed_algebra"
- shows "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"
- by (rule bounded_linear.sums [OF bounded_linear_mult_right])
-
-lemma summable_mult:
- fixes c :: "'a::real_normed_algebra"
- shows "summable f \<Longrightarrow> summable (%n. c * f n)"
- by (rule bounded_linear.summable [OF bounded_linear_mult_right])
-
-lemma suminf_mult:
- fixes c :: "'a::real_normed_algebra"
- shows "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f"
- by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric])
-
-lemma sums_mult2:
- fixes c :: "'a::real_normed_algebra"
- shows "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"
- by (rule bounded_linear.sums [OF bounded_linear_mult_left])
-
-lemma summable_mult2:
- fixes c :: "'a::real_normed_algebra"
- shows "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"
- by (rule bounded_linear.summable [OF bounded_linear_mult_left])
-
-lemma suminf_mult2:
- fixes c :: "'a::real_normed_algebra"
- shows "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"
- by (rule bounded_linear.suminf [OF bounded_linear_mult_left])
-
-lemma sums_divide:
- fixes c :: "'a::real_normed_field"
- shows "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"
- by (rule bounded_linear.sums [OF bounded_linear_divide])
-
-lemma summable_divide:
- fixes c :: "'a::real_normed_field"
- shows "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"
- by (rule bounded_linear.summable [OF bounded_linear_divide])
-
-lemma suminf_divide:
- fixes c :: "'a::real_normed_field"
- shows "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"
- by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric])
-
-lemma sums_add:
- fixes a b :: "'a::real_normed_field"
- shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n + Y n) sums (a + b)"
- unfolding sums_def by (simp add: setsum_addf tendsto_add)
-
-lemma summable_add:
- fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n + Y n)"
-unfolding summable_def by (auto intro: sums_add)
+ by (rule sums_zero [THEN sums_unique, symmetric])
-lemma suminf_add:
- fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X + suminf Y = (\<Sum>n. X n + Y n)"
-by (intro sums_unique sums_add summable_sums)
-
-lemma sums_diff:
- fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>X sums a; Y sums b\<rbrakk> \<Longrightarrow> (\<lambda>n. X n - Y n) sums (a - b)"
- unfolding sums_def by (simp add: setsum_subtractf tendsto_diff)
-
-lemma summable_diff:
- fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> summable (\<lambda>n. X n - Y n)"
-unfolding summable_def by (auto intro: sums_diff)
-
-lemma suminf_diff:
- fixes X Y :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>summable X; summable Y\<rbrakk> \<Longrightarrow> suminf X - suminf Y = (\<Sum>n. X n - Y n)"
-by (intro sums_unique sums_diff summable_sums)
-
-lemma sums_minus:
- fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "X sums a ==> (\<lambda>n. - X n) sums (- a)"
- unfolding sums_def by (simp add: setsum_negf tendsto_minus)
-
-lemma summable_minus:
- fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "summable X \<Longrightarrow> summable (\<lambda>n. - X n)"
-unfolding summable_def by (auto intro: sums_minus)
-
-lemma suminf_minus:
- fixes X :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "summable X \<Longrightarrow> (\<Sum>n. - X n) = - (\<Sum>n. X n)"
-by (intro sums_unique [symmetric] sums_minus summable_sums)
+context
+ fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
+begin
-lemma sums_group:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_field"
- shows "\<lbrakk>f sums s; 0 < k\<rbrakk> \<Longrightarrow> (\<lambda>n. setsum f {n*k..<n*k+k}) sums s"
-apply (simp only: sums_def sumr_group)
-apply (unfold LIMSEQ_iff, safe)
-apply (drule_tac x="r" in spec, safe)
-apply (rule_tac x="no" in exI, safe)
-apply (drule_tac x="n*k" in spec)
-apply (erule mp)
-apply (erule order_trans)
-apply simp
-done
-
-text{*A summable series of positive terms has limit that is at least as
-great as any partial sum.*}
-
-lemma pos_summable:
- fixes f:: "nat \<Rightarrow> real"
- assumes pos: "\<And>n. 0 \<le> f n" and le: "\<And>n. setsum f {0..<n} \<le> x"
- shows "summable f"
-proof -
- have "convergent (\<lambda>n. setsum f {0..<n})"
- proof (rule Bseq_mono_convergent)
- show "Bseq (\<lambda>n. setsum f {0..<n})"
- by (intro BseqI'[of _ x]) (auto simp add: setsum_nonneg pos intro: le)
- next
- show "\<forall>m n. m \<le> n \<longrightarrow> setsum f {0..<m} \<le> setsum f {0..<n}"
- by (auto intro: setsum_mono2 pos)
- qed
- thus ?thesis
- by (force simp add: summable_def sums_def convergent_def)
-qed
-
-lemma series_pos_le:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 \<le> f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} \<le> suminf f"
- apply (drule summable_sums)
- apply (simp add: sums_def)
- apply (rule LIMSEQ_le_const)
+lemma series_pos_le: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> setsum f {..<n} \<le> suminf f"
+ apply (rule LIMSEQ_le_const[OF summable_LIMSEQ])
apply assumption
apply (intro exI[of _ n])
- apply (auto intro!: setsum_mono2)
+ apply (auto intro!: setsum_mono2 simp: not_le[symmetric])
done
-lemma series_pos_less:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_ab_semigroup_add_imp_le, ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {0..<n} < suminf f"
- apply (rule_tac y="setsum f {0..<Suc n}" in order_less_le_trans)
- using add_less_cancel_left [of "setsum f {0..<n}" 0 "f n"]
- apply simp
- apply (erule series_pos_le)
- apply (simp add: order_less_imp_le)
- done
-
-lemma suminf_eq_zero_iff:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"
+lemma suminf_eq_zero_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"
proof
assume "summable f" "suminf f = 0" and pos: "\<forall>n. 0 \<le> f n"
then have "f sums 0"
@@ -419,77 +216,208 @@
qed
with pos show "\<forall>n. f n = 0"
by (auto intro!: antisym)
-next
- assume "\<forall>n. f n = 0"
- then have "f = (\<lambda>n. 0)"
- by auto
- then show "suminf f = 0"
- by simp
+qed (metis suminf_zero fun_eq_iff)
+
+lemma suminf_gt_zero_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"
+ using series_pos_le[of 0] suminf_eq_zero_iff by (simp add: less_le)
+
+lemma suminf_gt_zero: "summable f \<Longrightarrow> \<forall>n. 0 < f n \<Longrightarrow> 0 < suminf f"
+ using suminf_gt_zero_iff by (simp add: less_imp_le)
+
+lemma suminf_ge_zero: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 \<le> suminf f"
+ by (drule_tac n="0" in series_pos_le) simp_all
+
+lemma suminf_le: "summable f \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
+ by (metis LIMSEQ_le_const2 summable_LIMSEQ)
+
+lemma summable_le: "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"
+ by (rule LIMSEQ_le) (auto intro: setsum_mono summable_LIMSEQ)
+
+end
+
+lemma series_pos_less:
+ fixes f :: "nat \<Rightarrow> 'a::{ordered_ab_semigroup_add_imp_le, ordered_comm_monoid_add, linorder_topology}"
+ shows "\<lbrakk>summable f; \<forall>m\<ge>n. 0 < f m\<rbrakk> \<Longrightarrow> setsum f {..<n} < suminf f"
+ apply simp
+ apply (rule_tac y="setsum f {..<Suc n}" in order_less_le_trans)
+ using add_less_cancel_left [of "setsum f {..<n}" 0 "f n"]
+ apply simp
+ apply (erule series_pos_le)
+ apply (simp add: order_less_imp_le)
+ done
+
+lemma sums_Suc_iff:
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+ shows "(\<lambda>n. f (Suc n)) sums s \<longleftrightarrow> f sums (s + f 0)"
+proof -
+ have "f sums (s + f 0) \<longleftrightarrow> (\<lambda>i. \<Sum>j<Suc i. f j) ----> s + f 0"
+ by (subst LIMSEQ_Suc_iff) (simp add: sums_def)
+ also have "\<dots> \<longleftrightarrow> (\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"
+ by (simp add: ac_simps setsum_reindex image_iff lessThan_Suc_eq_insert_0)
+ also have "\<dots> \<longleftrightarrow> (\<lambda>n. f (Suc n)) sums s"
+ proof
+ assume "(\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"
+ with tendsto_add[OF this tendsto_const, of "- f 0"]
+ show "(\<lambda>i. f (Suc i)) sums s"
+ by (simp add: sums_def)
+ qed (auto intro: tendsto_add tendsto_const simp: sums_def)
+ finally show ?thesis ..
qed
-lemma suminf_gt_zero_iff:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"
- using series_pos_le[of f 0] suminf_eq_zero_iff[of f]
- by (simp add: less_le)
+context
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
+begin
+
+lemma sums_add: "f sums a \<Longrightarrow> g sums b \<Longrightarrow> (\<lambda>n. f n + g n) sums (a + b)"
+ unfolding sums_def by (simp add: setsum_addf tendsto_add)
+
+lemma summable_add: "summable f \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. f n + g n)"
+ unfolding summable_def by (auto intro: sums_add)
+
+lemma suminf_add: "summable f \<Longrightarrow> summable g \<Longrightarrow> suminf f + suminf g = (\<Sum>n. f n + g n)"
+ by (intro sums_unique sums_add summable_sums)
+
+lemma sums_diff: "f sums a \<Longrightarrow> g sums b \<Longrightarrow> (\<lambda>n. f n - g n) sums (a - b)"
+ unfolding sums_def by (simp add: setsum_subtractf tendsto_diff)
+
+lemma summable_diff: "summable f \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. f n - g n)"
+ unfolding summable_def by (auto intro: sums_diff)
+
+lemma suminf_diff: "summable f \<Longrightarrow> summable g \<Longrightarrow> suminf f - suminf g = (\<Sum>n. f n - g n)"
+ by (intro sums_unique sums_diff summable_sums)
+
+lemma sums_minus: "f sums a \<Longrightarrow> (\<lambda>n. - f n) sums (- a)"
+ unfolding sums_def by (simp add: setsum_negf tendsto_minus)
+
+lemma summable_minus: "summable f \<Longrightarrow> summable (\<lambda>n. - f n)"
+ unfolding summable_def by (auto intro: sums_minus)
-lemma suminf_gt_zero:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>n. 0 < f n\<rbrakk> \<Longrightarrow> 0 < suminf f"
- using suminf_gt_zero_iff[of f] by (simp add: less_imp_le)
+lemma suminf_minus: "summable f \<Longrightarrow> (\<Sum>n. - f n) = - (\<Sum>n. f n)"
+ by (intro sums_unique [symmetric] sums_minus summable_sums)
+
+lemma sums_Suc: "(\<lambda> n. f (Suc n)) sums l \<Longrightarrow> f sums (l + f 0)"
+ by (simp add: sums_Suc_iff)
+
+lemma sums_iff_shift: "(\<lambda>i. f (i + n)) sums s \<longleftrightarrow> f sums (s + (\<Sum>i<n. f i))"
+proof (induct n arbitrary: s)
+ case (Suc n)
+ moreover have "(\<lambda>i. f (Suc i + n)) sums s \<longleftrightarrow> (\<lambda>i. f (i + n)) sums (s + f n)"
+ by (subst sums_Suc_iff) simp
+ ultimately show ?case
+ by (simp add: ac_simps)
+qed simp
-lemma suminf_ge_zero:
- fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>summable f; \<forall>n. 0 \<le> f n\<rbrakk> \<Longrightarrow> 0 \<le> suminf f"
- by (drule_tac n="0" in series_pos_le, simp_all)
+lemma summable_iff_shift: "summable (\<lambda>n. f (n + k)) \<longleftrightarrow> summable f"
+ by (metis diff_add_cancel summable_def sums_iff_shift[abs_def])
+
+lemma sums_split_initial_segment: "f sums s \<Longrightarrow> (\<lambda>i. f (i + n)) sums (s - (\<Sum>i<n. f i))"
+ by (simp add: sums_iff_shift)
+
+lemma summable_ignore_initial_segment: "summable f \<Longrightarrow> summable (\<lambda>n. f(n + k))"
+ by (simp add: summable_iff_shift)
+
+lemma suminf_minus_initial_segment: "summable f \<Longrightarrow> (\<Sum>n. f (n + k)) = (\<Sum>n. f n) - (\<Sum>i<k. f i)"
+ by (rule sums_unique[symmetric]) (auto simp: sums_iff_shift)
+
+lemma suminf_split_initial_segment: "summable f \<Longrightarrow> suminf f = (\<Sum>n. f(n + k)) + (\<Sum>i<k. f i)"
+ by (auto simp add: suminf_minus_initial_segment)
-lemma sumr_pos_lt_pair:
- fixes f :: "nat \<Rightarrow> real"
- shows "\<lbrakk>summable f;
- \<And>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>
- \<Longrightarrow> setsum f {0..<k} < suminf f"
-unfolding One_nat_def
-apply (subst suminf_split_initial_segment [where k="k"])
-apply assumption
-apply simp
-apply (drule_tac k="k" in summable_ignore_initial_segment)
-apply (drule_tac k="Suc (Suc 0)" in sums_group [OF summable_sums], simp)
-apply simp
-apply (frule sums_unique)
-apply (drule sums_summable)
-apply simp
-apply (erule suminf_gt_zero)
-apply (simp add: add_ac)
-done
+lemma suminf_exist_split:
+ fixes r :: real assumes "0 < r" and "summable f"
+ shows "\<exists>N. \<forall>n\<ge>N. norm (\<Sum>i. f (i + n)) < r"
+proof -
+ from LIMSEQ_D[OF summable_LIMSEQ[OF `summable f`] `0 < r`]
+ obtain N :: nat where "\<forall> n \<ge> N. norm (setsum f {..<n} - suminf f) < r" by auto
+ thus ?thesis
+ by (auto simp: norm_minus_commute suminf_minus_initial_segment[OF `summable f`])
+qed
+
+lemma summable_LIMSEQ_zero: "summable f \<Longrightarrow> f ----> 0"
+ apply (drule summable_iff_convergent [THEN iffD1])
+ apply (drule convergent_Cauchy)
+ apply (simp only: Cauchy_iff LIMSEQ_iff, safe)
+ apply (drule_tac x="r" in spec, safe)
+ apply (rule_tac x="M" in exI, safe)
+ apply (drule_tac x="Suc n" in spec, simp)
+ apply (drule_tac x="n" in spec, simp)
+ done
+
+end
+
+lemma (in bounded_linear) sums: "(\<lambda>n. X n) sums a \<Longrightarrow> (\<lambda>n. f (X n)) sums (f a)"
+ unfolding sums_def by (drule tendsto, simp only: setsum)
+
+lemma (in bounded_linear) summable: "summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"
+ unfolding summable_def by (auto intro: sums)
+
+lemma (in bounded_linear) suminf: "summable (\<lambda>n. X n) \<Longrightarrow> f (\<Sum>n. X n) = (\<Sum>n. f (X n))"
+ by (intro sums_unique sums summable_sums)
+
+lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real]
+lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real]
+lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real]
+
+context
+ fixes f :: "nat \<Rightarrow> 'a::real_normed_algebra"
+begin
+
+lemma sums_mult: "f sums a \<Longrightarrow> (\<lambda>n. c * f n) sums (c * a)"
+ by (rule bounded_linear.sums [OF bounded_linear_mult_right])
+
+lemma summable_mult: "summable f \<Longrightarrow> summable (\<lambda>n. c * f n)"
+ by (rule bounded_linear.summable [OF bounded_linear_mult_right])
+
+lemma suminf_mult: "summable f \<Longrightarrow> suminf (\<lambda>n. c * f n) = c * suminf f"
+ by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric])
+
+lemma sums_mult2: "f sums a \<Longrightarrow> (\<lambda>n. f n * c) sums (a * c)"
+ by (rule bounded_linear.sums [OF bounded_linear_mult_left])
+
+lemma summable_mult2: "summable f \<Longrightarrow> summable (\<lambda>n. f n * c)"
+ by (rule bounded_linear.summable [OF bounded_linear_mult_left])
+
+lemma suminf_mult2: "summable f \<Longrightarrow> suminf f * c = (\<Sum>n. f n * c)"
+ by (rule bounded_linear.suminf [OF bounded_linear_mult_left])
+
+end
+
+context
+ fixes c :: "'a::real_normed_field"
+begin
+
+lemma sums_divide: "f sums a \<Longrightarrow> (\<lambda>n. f n / c) sums (a / c)"
+ by (rule bounded_linear.sums [OF bounded_linear_divide])
+
+lemma summable_divide: "summable f \<Longrightarrow> summable (\<lambda>n. f n / c)"
+ by (rule bounded_linear.summable [OF bounded_linear_divide])
+
+lemma suminf_divide: "summable f \<Longrightarrow> suminf (\<lambda>n. f n / c) = suminf f / c"
+ by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric])
text{*Sum of a geometric progression.*}
-lemmas sumr_geometric = geometric_sum [where 'a = real]
-
-lemma geometric_sums:
- fixes x :: "'a::{real_normed_field}"
- shows "norm x < 1 \<Longrightarrow> (\<lambda>n. x ^ n) sums (1 / (1 - x))"
+lemma geometric_sums: "norm c < 1 \<Longrightarrow> (\<lambda>n. c^n) sums (1 / (1 - c))"
proof -
- assume less_1: "norm x < 1"
- hence neq_1: "x \<noteq> 1" by auto
- hence neq_0: "x - 1 \<noteq> 0" by simp
- from less_1 have lim_0: "(\<lambda>n. x ^ n) ----> 0"
+ assume less_1: "norm c < 1"
+ hence neq_1: "c \<noteq> 1" by auto
+ hence neq_0: "c - 1 \<noteq> 0" by simp
+ from less_1 have lim_0: "(\<lambda>n. c^n) ----> 0"
by (rule LIMSEQ_power_zero)
- hence "(\<lambda>n. x ^ n / (x - 1) - 1 / (x - 1)) ----> 0 / (x - 1) - 1 / (x - 1)"
+ hence "(\<lambda>n. c ^ n / (c - 1) - 1 / (c - 1)) ----> 0 / (c - 1) - 1 / (c - 1)"
using neq_0 by (intro tendsto_intros)
- hence "(\<lambda>n. (x ^ n - 1) / (x - 1)) ----> 1 / (1 - x)"
+ hence "(\<lambda>n. (c ^ n - 1) / (c - 1)) ----> 1 / (1 - c)"
by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib)
- thus "(\<lambda>n. x ^ n) sums (1 / (1 - x))"
+ thus "(\<lambda>n. c ^ n) sums (1 / (1 - c))"
by (simp add: sums_def geometric_sum neq_1)
qed
-lemma summable_geometric:
- fixes x :: "'a::{real_normed_field}"
- shows "norm x < 1 \<Longrightarrow> summable (\<lambda>n. x ^ n)"
-by (rule geometric_sums [THEN sums_summable])
+lemma summable_geometric: "norm c < 1 \<Longrightarrow> summable (\<lambda>n. c^n)"
+ by (rule geometric_sums [THEN sums_summable])
-lemma half: "0 < 1 / (2::'a::linordered_field)"
- by simp
+lemma suminf_geometric: "norm c < 1 \<Longrightarrow> suminf (\<lambda>n. c^n) = 1 / (1 - c)"
+ by (rule sums_unique[symmetric]) (rule geometric_sums)
+
+end
lemma power_half_series: "(\<lambda>n. (1/2::real)^Suc n) sums 1"
proof -
@@ -503,110 +431,104 @@
text{*Cauchy-type criterion for convergence of series (c.f. Harrison)*}
-lemma summable_convergent_sumr_iff:
- "summable f = convergent (%n. setsum f {0..<n})"
-by (simp add: summable_def sums_def convergent_def)
+lemma summable_Cauchy:
+ fixes f :: "nat \<Rightarrow> 'a::banach"
+ shows "summable f \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n. norm (setsum f {m..<n}) < e)"
+ apply (simp only: summable_iff_convergent Cauchy_convergent_iff [symmetric] Cauchy_iff, safe)
+ apply (drule spec, drule (1) mp)
+ apply (erule exE, rule_tac x="M" in exI, clarify)
+ apply (rule_tac x="m" and y="n" in linorder_le_cases)
+ apply (frule (1) order_trans)
+ apply (drule_tac x="n" in spec, drule (1) mp)
+ apply (drule_tac x="m" in spec, drule (1) mp)
+ apply (simp_all add: setsum_diff [symmetric])
+ apply (drule spec, drule (1) mp)
+ apply (erule exE, rule_tac x="N" in exI, clarify)
+ apply (rule_tac x="m" and y="n" in linorder_le_cases)
+ apply (subst norm_minus_commute)
+ apply (simp_all add: setsum_diff [symmetric])
+ done
-lemma summable_LIMSEQ_zero:
- fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"
- shows "summable f \<Longrightarrow> f ----> 0"
-apply (drule summable_convergent_sumr_iff [THEN iffD1])
-apply (drule convergent_Cauchy)
-apply (simp only: Cauchy_iff LIMSEQ_iff, safe)
-apply (drule_tac x="r" in spec, safe)
-apply (rule_tac x="M" in exI, safe)
-apply (drule_tac x="Suc n" in spec, simp)
-apply (drule_tac x="n" in spec, simp)
-done
+context
+ fixes f :: "nat \<Rightarrow> 'a::banach"
+begin
+
+text{*Absolute convergence imples normal convergence*}
-lemma suminf_le:
- fixes x :: "'a :: {ordered_comm_monoid_add, linorder_topology}"
- shows "summable f \<Longrightarrow> (!!n. setsum f {0..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
- apply (drule summable_sums)
- apply (simp add: sums_def)
- apply (rule LIMSEQ_le_const2)
- apply assumption
- apply auto
+lemma summable_norm_cancel:
+ "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"
+ apply (simp only: summable_Cauchy, safe)
+ apply (drule_tac x="e" in spec, safe)
+ apply (rule_tac x="N" in exI, safe)
+ apply (drule_tac x="m" in spec, safe)
+ apply (rule order_le_less_trans [OF norm_setsum])
+ apply (rule order_le_less_trans [OF abs_ge_self])
+ apply simp
done
-lemma summable_Cauchy:
- "summable (f::nat \<Rightarrow> 'a::banach) =
- (\<forall>e > 0. \<exists>N. \<forall>m \<ge> N. \<forall>n. norm (setsum f {m..<n}) < e)"
-apply (simp only: summable_convergent_sumr_iff Cauchy_convergent_iff [symmetric] Cauchy_iff, safe)
-apply (drule spec, drule (1) mp)
-apply (erule exE, rule_tac x="M" in exI, clarify)
-apply (rule_tac x="m" and y="n" in linorder_le_cases)
-apply (frule (1) order_trans)
-apply (drule_tac x="n" in spec, drule (1) mp)
-apply (drule_tac x="m" in spec, drule (1) mp)
-apply (simp add: setsum_diff [symmetric])
-apply simp
-apply (drule spec, drule (1) mp)
-apply (erule exE, rule_tac x="N" in exI, clarify)
-apply (rule_tac x="m" and y="n" in linorder_le_cases)
-apply (subst norm_minus_commute)
-apply (simp add: setsum_diff [symmetric])
-apply (simp add: setsum_diff [symmetric])
-done
+lemma summable_norm: "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"
+ by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_setsum)
+
+text {* Comparison tests *}
-text{*Comparison test*}
+lemma summable_comparison_test: "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f"
+ apply (simp add: summable_Cauchy, safe)
+ apply (drule_tac x="e" in spec, safe)
+ apply (rule_tac x = "N + Na" in exI, safe)
+ apply (rotate_tac 2)
+ apply (drule_tac x = m in spec)
+ apply (auto, rotate_tac 2, drule_tac x = n in spec)
+ apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)
+ apply (rule norm_setsum)
+ apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
+ apply (auto intro: setsum_mono simp add: abs_less_iff)
+ done
+
+subsection {* The Ratio Test*}
-lemma norm_setsum:
- fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
- shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"
-apply (case_tac "finite A")
-apply (erule finite_induct)
-apply simp
-apply simp
-apply (erule order_trans [OF norm_triangle_ineq add_left_mono])
-apply simp
-done
-
-lemma norm_bound_subset:
- fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
- assumes "finite s" "t \<subseteq> s"
- assumes le: "(\<And>x. x \<in> s \<Longrightarrow> norm(f x) \<le> g x)"
- shows "norm (setsum f t) \<le> setsum g s"
-proof -
- have "norm (setsum f t) \<le> (\<Sum>i\<in>t. norm (f i))"
- by (rule norm_setsum)
- also have "\<dots> \<le> (\<Sum>i\<in>t. g i)"
- using assms by (auto intro!: setsum_mono)
- also have "\<dots> \<le> setsum g s"
- using assms order.trans[OF norm_ge_zero le]
- by (auto intro!: setsum_mono3)
- finally show ?thesis .
+lemma summable_ratio_test:
+ assumes "c < 1" "\<And>n. n \<ge> N \<Longrightarrow> norm (f (Suc n)) \<le> c * norm (f n)"
+ shows "summable f"
+proof cases
+ assume "0 < c"
+ show "summable f"
+ proof (rule summable_comparison_test)
+ show "\<exists>N'. \<forall>n\<ge>N'. norm (f n) \<le> (norm (f N) / (c ^ N)) * c ^ n"
+ proof (intro exI allI impI)
+ fix n assume "N \<le> n" then show "norm (f n) \<le> (norm (f N) / (c ^ N)) * c ^ n"
+ proof (induct rule: inc_induct)
+ case (step m)
+ moreover have "norm (f (Suc m)) / c ^ Suc m * c ^ n \<le> norm (f m) / c ^ m * c ^ n"
+ using `0 < c` `c < 1` assms(2)[OF `N \<le> m`] by (simp add: field_simps)
+ ultimately show ?case by simp
+ qed (insert `0 < c`, simp)
+ qed
+ show "summable (\<lambda>n. norm (f N) / c ^ N * c ^ n)"
+ using `0 < c` `c < 1` by (intro summable_mult summable_geometric) simp
+ qed
+next
+ assume c: "\<not> 0 < c"
+ { fix n assume "n \<ge> N"
+ then have "norm (f (Suc n)) \<le> c * norm (f n)"
+ by fact
+ also have "\<dots> \<le> 0"
+ using c by (simp add: not_less mult_nonpos_nonneg)
+ finally have "f (Suc n) = 0"
+ by auto }
+ then show "summable f"
+ by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_iff)
qed
-lemma summable_comparison_test:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f"
-apply (simp add: summable_Cauchy, safe)
-apply (drule_tac x="e" in spec, safe)
-apply (rule_tac x = "N + Na" in exI, safe)
-apply (rotate_tac 2)
-apply (drule_tac x = m in spec)
-apply (auto, rotate_tac 2, drule_tac x = n in spec)
-apply (rule_tac y = "\<Sum>k=m..<n. norm (f k)" in order_le_less_trans)
-apply (rule norm_setsum)
-apply (rule_tac y = "setsum g {m..<n}" in order_le_less_trans)
-apply (auto intro: setsum_mono simp add: abs_less_iff)
-done
+end
lemma summable_norm_comparison_test:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk>
- \<Longrightarrow> summable (\<lambda>n. norm (f n))"
-apply (rule summable_comparison_test)
-apply (auto)
-done
+ "\<lbrakk>\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable (\<lambda>n. norm (f n))"
+ by (rule summable_comparison_test) auto
-lemma summable_rabs_comparison_test:
+lemma summable_rabs_cancel:
fixes f :: "nat \<Rightarrow> real"
- shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable (\<lambda>n. \<bar>f n\<bar>)"
-apply (rule summable_comparison_test)
-apply (auto)
-done
+ shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"
+ by (rule summable_norm_cancel) simp
text{*Summability of geometric series for real algebras*}
@@ -620,119 +542,34 @@
by (simp add: summable_geometric)
qed
-text{*Limit comparison property for series (c.f. jrh)*}
-lemma summable_le:
- fixes f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "\<lbrakk>\<forall>n. f n \<le> g n; summable f; summable g\<rbrakk> \<Longrightarrow> suminf f \<le> suminf g"
-apply (drule summable_sums)+
-apply (simp only: sums_def, erule (1) LIMSEQ_le)
-apply (rule exI)
-apply (auto intro!: setsum_mono)
-done
-
-lemma summable_le2:
- fixes f g :: "nat \<Rightarrow> real"
- shows "\<lbrakk>\<forall>n. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable f \<and> suminf f \<le> suminf g"
-apply (subgoal_tac "summable f")
-apply (auto intro!: summable_le)
-apply (simp add: abs_le_iff)
-apply (rule_tac g="g" in summable_comparison_test, simp_all)
-done
+text{*A summable series of positive terms has limit that is at least as
+great as any partial sum.*}
-(* specialisation for the common 0 case *)
-lemma suminf_0_le:
- fixes f::"nat\<Rightarrow>real"
- assumes gt0: "\<forall>n. 0 \<le> f n" and sm: "summable f"
- shows "0 \<le> suminf f"
- using suminf_ge_zero[OF sm gt0] by simp
+lemma pos_summable:
+ fixes f:: "nat \<Rightarrow> real"
+ assumes pos: "\<And>n. 0 \<le> f n" and le: "\<And>n. setsum f {..<n} \<le> x"
+ shows "summable f"
+proof -
+ have "convergent (\<lambda>n. setsum f {..<n})"
+ proof (rule Bseq_mono_convergent)
+ show "Bseq (\<lambda>n. setsum f {..<n})"
+ by (intro BseqI'[of _ x]) (auto simp add: setsum_nonneg pos intro: le)
+ qed (auto intro: setsum_mono2 pos)
+ thus ?thesis
+ by (force simp add: summable_def sums_def convergent_def)
+qed
-text{*Absolute convergence imples normal convergence*}
-lemma summable_norm_cancel:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> summable f"
-apply (simp only: summable_Cauchy, safe)
-apply (drule_tac x="e" in spec, safe)
-apply (rule_tac x="N" in exI, safe)
-apply (drule_tac x="m" in spec, safe)
-apply (rule order_le_less_trans [OF norm_setsum])
-apply (rule order_le_less_trans [OF abs_ge_self])
-apply simp
-done
-
-lemma summable_rabs_cancel:
+lemma summable_rabs_comparison_test:
fixes f :: "nat \<Rightarrow> real"
- shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> summable f"
-by (rule summable_norm_cancel, simp)
-
-text{*Absolute convergence of series*}
-lemma summable_norm:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "summable (\<lambda>n. norm (f n)) \<Longrightarrow> norm (suminf f) \<le> (\<Sum>n. norm (f n))"
- by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel
- summable_sumr_LIMSEQ_suminf norm_setsum)
+ shows "\<lbrakk>\<exists>N. \<forall>n\<ge>N. \<bar>f n\<bar> \<le> g n; summable g\<rbrakk> \<Longrightarrow> summable (\<lambda>n. \<bar>f n\<bar>)"
+ by (rule summable_comparison_test) auto
lemma summable_rabs:
fixes f :: "nat \<Rightarrow> real"
shows "summable (\<lambda>n. \<bar>f n\<bar>) \<Longrightarrow> \<bar>suminf f\<bar> \<le> (\<Sum>n. \<bar>f n\<bar>)"
by (fold real_norm_def, rule summable_norm)
-subsection{* The Ratio Test*}
-
-lemma norm_ratiotest_lemma:
- fixes x y :: "'a::real_normed_vector"
- shows "\<lbrakk>c \<le> 0; norm x \<le> c * norm y\<rbrakk> \<Longrightarrow> x = 0"
-apply (subgoal_tac "norm x \<le> 0", simp)
-apply (erule order_trans)
-apply (simp add: mult_le_0_iff)
-done
-
-lemma rabs_ratiotest_lemma: "[| c \<le> 0; abs x \<le> c * abs y |] ==> x = (0::real)"
-by (erule norm_ratiotest_lemma, simp)
-
-(* TODO: MOVE *)
-lemma le_Suc_ex: "(k::nat) \<le> l ==> (\<exists>n. l = k + n)"
-apply (drule le_imp_less_or_eq)
-apply (auto dest: less_imp_Suc_add)
-done
-
-lemma le_Suc_ex_iff: "((k::nat) \<le> l) = (\<exists>n. l = k + n)"
-by (auto simp add: le_Suc_ex)
-
-(*All this trouble just to get 0<c *)
-lemma ratio_test_lemma2:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "\<lbrakk>\<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> 0 < c \<or> summable f"
-apply (simp (no_asm) add: linorder_not_le [symmetric])
-apply (simp add: summable_Cauchy)
-apply (safe, subgoal_tac "\<forall>n. N < n --> f (n) = 0")
- prefer 2
- apply clarify
- apply(erule_tac x = "n - Suc 0" in allE)
- apply (simp add:diff_Suc split:nat.splits)
- apply (blast intro: norm_ratiotest_lemma)
-apply (rule_tac x = "Suc N" in exI, clarify)
-apply(simp cong del: setsum_cong cong: setsum_ivl_cong)
-done
-
-lemma ratio_test:
- fixes f :: "nat \<Rightarrow> 'a::banach"
- shows "\<lbrakk>c < 1; \<forall>n\<ge>N. norm (f (Suc n)) \<le> c * norm (f n)\<rbrakk> \<Longrightarrow> summable f"
-apply (frule ratio_test_lemma2, auto)
-apply (rule_tac g = "%n. (norm (f N) / (c ^ N))*c ^ n"
- in summable_comparison_test)
-apply (rule_tac x = N in exI, safe)
-apply (drule le_Suc_ex_iff [THEN iffD1])
-apply (auto simp add: power_add field_power_not_zero)
-apply (induct_tac "na", auto)
-apply (rule_tac y = "c * norm (f (N + n))" in order_trans)
-apply (auto intro: mult_right_mono simp add: summable_def)
-apply (rule_tac x = "norm (f N) * (1/ (1 - c)) / (c ^ N)" in exI)
-apply (rule sums_divide)
-apply (rule sums_mult)
-apply (auto intro!: geometric_sums)
-done
-
subsection {* Cauchy Product Formula *}
text {*
@@ -742,14 +579,14 @@
lemma setsum_triangle_reindex:
fixes n :: nat
- shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k=0..<n. \<Sum>i=0..k. f i (k - i))"
+ shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k<n. \<Sum>i=0..k. f i (k - i))"
proof -
have "(\<Sum>(i, j)\<in>{(i, j). i + j < n}. f i j) =
- (\<Sum>(k, i)\<in>(SIGMA k:{0..<n}. {0..k}). f i (k - i))"
+ (\<Sum>(k, i)\<in>(SIGMA k:{..<n}. {0..k}). f i (k - i))"
proof (rule setsum_reindex_cong)
- show "inj_on (\<lambda>(k,i). (i, k - i)) (SIGMA k:{0..<n}. {0..k})"
+ show "inj_on (\<lambda>(k,i). (i, k - i)) (SIGMA k:{..<n}. {0..k})"
by (rule inj_on_inverseI [where g="\<lambda>(i,j). (i+j, i)"], auto)
- show "{(i,j). i + j < n} = (\<lambda>(k,i). (i, k - i)) ` (SIGMA k:{0..<n}. {0..k})"
+ show "{(i,j). i + j < n} = (\<lambda>(k,i). (i, k - i)) ` (SIGMA k:{..<n}. {0..k})"
by (safe, rule_tac x="(a+b,a)" in image_eqI, auto)
show "\<And>a. (\<lambda>(k, i). f i (k - i)) a = split f ((\<lambda>(k, i). (i, k - i)) a)"
by clarify
@@ -763,7 +600,7 @@
assumes b: "summable (\<lambda>k. norm (b k))"
shows "(\<lambda>k. \<Sum>i=0..k. a i * b (k - i)) sums ((\<Sum>k. a k) * (\<Sum>k. b k))"
proof -
- let ?S1 = "\<lambda>n::nat. {0..<n} \<times> {0..<n}"
+ let ?S1 = "\<lambda>n::nat. {..<n} \<times> {..<n}"
let ?S2 = "\<lambda>n::nat. {(i,j). i + j < n}"
have S1_mono: "\<And>m n. m \<le> n \<Longrightarrow> ?S1 m \<subseteq> ?S1 n" by auto
have S2_le_S1: "\<And>n. ?S2 n \<subseteq> ?S1 n" by auto
@@ -779,20 +616,15 @@
unfolding real_norm_def
by (simp only: abs_of_nonneg setsum_nonneg [rule_format])
- have "(\<lambda>n. (\<Sum>k=0..<n. a k) * (\<Sum>k=0..<n. b k))
- ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
- by (intro tendsto_mult summable_sumr_LIMSEQ_suminf
- summable_norm_cancel [OF a] summable_norm_cancel [OF b])
+ have "(\<lambda>n. (\<Sum>k<n. a k) * (\<Sum>k<n. b k)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
+ by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b])
hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
- by (simp only: setsum_product setsum_Sigma [rule_format]
- finite_atLeastLessThan)
+ by (simp only: setsum_product setsum_Sigma [rule_format] finite_lessThan)
- have "(\<lambda>n. (\<Sum>k=0..<n. norm (a k)) * (\<Sum>k=0..<n. norm (b k)))
- ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
- using a b by (intro tendsto_mult summable_sumr_LIMSEQ_suminf)
+ have "(\<lambda>n. (\<Sum>k<n. norm (a k)) * (\<Sum>k<n. norm (b k))) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
+ using a b by (intro tendsto_mult summable_LIMSEQ)
hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
- by (simp only: setsum_product setsum_Sigma [rule_format]
- finite_atLeastLessThan)
+ by (simp only: setsum_product setsum_Sigma [rule_format] finite_lessThan)
hence "convergent (\<lambda>n. setsum ?f (?S1 n))"
by (rule convergentI)
hence Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"
--- a/src/HOL/Set_Interval.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Set_Interval.thy Tue Mar 18 15:53:48 2014 +0100
@@ -1472,10 +1472,10 @@
lemma geometric_sum:
assumes "x \<noteq> 1"
- shows "(\<Sum>i=0..<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::field)"
+ shows "(\<Sum>i<n. x ^ i) = (x ^ n - 1) / (x - 1::'a::field)"
proof -
from assms obtain y where "y = x - 1" and "y \<noteq> 0" by simp_all
- moreover have "(\<Sum>i=0..<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
+ moreover have "(\<Sum>i<n. (y + 1) ^ i) = ((y + 1) ^ n - 1) / y"
proof (induct n)
case 0 then show ?case by simp
next
@@ -1490,8 +1490,7 @@
subsection {* The formula for arithmetic sums *}
lemma gauss_sum:
- "(2::'a::comm_semiring_1)*(\<Sum>i\<in>{1..n}. of_nat i) =
- of_nat n*((of_nat n)+1)"
+ "(2::'a::comm_semiring_1)*(\<Sum>i\<in>{1..n}. of_nat i) = of_nat n*((of_nat n)+1)"
proof (induct n)
case 0
show ?case by simp
@@ -1575,8 +1574,8 @@
lemma nat_diff_setsum_reindex:
fixes x :: "'a::{comm_ring,monoid_mult}"
- shows "(\<Sum>i=0..<n. f (n - Suc i)) = (\<Sum>i=0..<n. f i)"
-apply (subst setsum_reindex_cong [of "%i. n - Suc i" "{0..< n}"])
+ shows "(\<Sum>i<n. f (n - Suc i)) = (\<Sum>i<n. f i)"
+apply (subst setsum_reindex_cong [of "%i. n - Suc i" "{..< n}"])
apply (auto simp: inj_on_def)
apply (rule_tac x="n - Suc x" in image_eqI, auto)
done
--- a/src/HOL/Taylor.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Taylor.thy Tue Mar 18 15:53:48 2014 +0100
@@ -18,8 +18,7 @@
and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
and INTERV: "a \<le> c" "c < b"
shows "\<exists> t. c < t & t < b &
- f b = setsum (%m. (diff m c / real (fact m)) * (b - c)^m) {0..<n} +
- (diff n t / real (fact n)) * (b - c)^n"
+ f b = (\<Sum>m<n. (diff m c / real (fact m)) * (b - c)^m) + (diff n t / real (fact n)) * (b - c)^n"
proof -
from INTERV have "0 < b-c" by arith
moreover
@@ -38,17 +37,17 @@
qed
ultimately
have EX:"EX t>0. t < b - c &
- f (b - c + c) = (SUM m = 0..<n. diff m (0 + c) / real (fact m) * (b - c) ^ m) +
+ f (b - c + c) = (SUM m<n. diff m (0 + c) / real (fact m) * (b - c) ^ m) +
diff n (t + c) / real (fact n) * (b - c) ^ n"
by (rule Maclaurin)
show ?thesis
proof -
from EX obtain x where
X: "0 < x & x < b - c &
- f (b - c + c) = (\<Sum>m = 0..<n. diff m (0 + c) / real (fact m) * (b - c) ^ m) +
+ f (b - c + c) = (\<Sum>m<n. diff m (0 + c) / real (fact m) * (b - c) ^ m) +
diff n (x + c) / real (fact n) * (b - c) ^ n" ..
let ?H = "x + c"
- from X have "c<?H & ?H<b \<and> f b = (\<Sum>m = 0..<n. diff m c / real (fact m) * (b - c) ^ m) +
+ from X have "c<?H & ?H<b \<and> f b = (\<Sum>m<n. diff m c / real (fact m) * (b - c) ^ m) +
diff n ?H / real (fact n) * (b - c) ^ n"
by fastforce
thus ?thesis by fastforce
@@ -60,8 +59,7 @@
and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
and INTERV: "a < c" "c \<le> b"
shows "\<exists> t. a < t & t < c &
- f a = setsum (% m. (diff m c / real (fact m)) * (a - c)^m) {0..<n} +
- (diff n t / real (fact n)) * (a - c)^n"
+ f a = (\<Sum>m<n. (diff m c / real (fact m)) * (a - c)^m) + (diff n t / real (fact n)) * (a - c)^n"
proof -
from INTERV have "a-c < 0" by arith
moreover
@@ -79,16 +77,16 @@
qed
ultimately
have EX: "EX t>a - c. t < 0 &
- f (a - c + c) = (SUM m = 0..<n. diff m (0 + c) / real (fact m) * (a - c) ^ m) +
+ f (a - c + c) = (SUM m<n. diff m (0 + c) / real (fact m) * (a - c) ^ m) +
diff n (t + c) / real (fact n) * (a - c) ^ n"
by (rule Maclaurin_minus)
show ?thesis
proof -
from EX obtain x where X: "a - c < x & x < 0 &
- f (a - c + c) = (SUM m = 0..<n. diff m (0 + c) / real (fact m) * (a - c) ^ m) +
+ f (a - c + c) = (SUM m<n. diff m (0 + c) / real (fact m) * (a - c) ^ m) +
diff n (x + c) / real (fact n) * (a - c) ^ n" ..
let ?H = "x + c"
- from X have "a<?H & ?H<c \<and> f a = (\<Sum>m = 0..<n. diff m c / real (fact m) * (a - c) ^ m) +
+ from X have "a<?H & ?H<c \<and> f a = (\<Sum>m<n. diff m c / real (fact m) * (a - c) ^ m) +
diff n ?H / real (fact n) * (a - c) ^ n"
by fastforce
thus ?thesis by fastforce
@@ -100,8 +98,7 @@
and DERIV: "(\<forall> m t. m < n & a \<le> t & t \<le> b \<longrightarrow> DERIV (diff m) t :> (diff (Suc m) t))"
and INTERV: "a \<le> c " "c \<le> b" "a \<le> x" "x \<le> b" "x \<noteq> c"
shows "\<exists> t. (if x<c then (x < t & t < c) else (c < t & t < x)) &
- f x = setsum (% m. (diff m c / real (fact m)) * (x - c)^m) {0..<n} +
- (diff n t / real (fact n)) * (x - c)^n"
+ f x = (\<Sum>m<n. (diff m c / real (fact m)) * (x - c)^m) + (diff n t / real (fact n)) * (x - c)^n"
proof (cases "x<c")
case True
note INIT
@@ -111,8 +108,7 @@
moreover note True
moreover from INTERV have "c \<le> b" by simp
ultimately have EX: "\<exists>t>x. t < c \<and> f x =
- (\<Sum>m = 0..<n. diff m c / real (fact m) * (x - c) ^ m) +
- diff n t / real (fact n) * (x - c) ^ n"
+ (\<Sum>m<n. diff m c / real (fact m) * (x - c) ^ m) + diff n t / real (fact n) * (x - c) ^ n"
by (rule taylor_down)
with True show ?thesis by simp
next
@@ -124,8 +120,7 @@
moreover from INTERV have "a \<le> c" by arith
moreover from False and INTERV have "c < x" by arith
ultimately have EX: "\<exists>t>c. t < x \<and> f x =
- (\<Sum>m = 0..<n. diff m c / real (fact m) * (x - c) ^ m) +
- diff n t / real (fact n) * (x - c) ^ n"
+ (\<Sum>m<n. diff m c / real (fact m) * (x - c) ^ m) + diff n t / real (fact n) * (x - c) ^ n"
by (rule taylor_up)
with False show ?thesis by simp
qed
--- a/src/HOL/Transcendental.thy Tue Mar 18 14:32:23 2014 +0100
+++ b/src/HOL/Transcendental.thy Tue Mar 18 15:53:48 2014 +0100
@@ -21,66 +21,55 @@
thus ?thesis by (simp add: power_commutes)
qed
-lemma lemma_realpow_diff_sumr:
- fixes y :: "'a::{comm_semiring_0,monoid_mult}"
- shows
- "(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) =
- y * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
- by (simp add: setsum_right_distrib lemma_realpow_diff mult_ac del: setsum_op_ivl_Suc)
-
lemma lemma_realpow_diff_sumr2:
fixes y :: "'a::{comm_ring,monoid_mult}"
shows
"x ^ (Suc n) - y ^ (Suc n) =
- (x - y) * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
+ (x - y) * (\<Sum>p<Suc n. (x ^ p) * y ^ (n - p))"
proof (induct n)
- case 0 show ?case
- by simp
-next
case (Suc n)
have "x ^ Suc (Suc n) - y ^ Suc (Suc n) = x * (x * x ^ n) - y * (y * y ^ n)"
by simp
also have "... = y * (x ^ (Suc n) - y ^ (Suc n)) + (x - y) * (x * x ^ n)"
by (simp add: algebra_simps)
- also have "... = y * ((x - y) * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))) + (x - y) * (x * x ^ n)"
+ also have "... = y * ((x - y) * (\<Sum>p<Suc n. (x ^ p) * y ^ (n - p))) + (x - y) * (x * x ^ n)"
by (simp only: Suc)
- also have "... = (x - y) * (y * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))) + (x - y) * (x * x ^ n)"
+ also have "... = (x - y) * (y * (\<Sum>p<Suc n. (x ^ p) * y ^ (n - p))) + (x - y) * (x * x ^ n)"
by (simp only: mult_left_commute)
- also have "... = (x - y) * (\<Sum>p = 0..<Suc (Suc n). x ^ p * y ^ (Suc n - p))"
- by (simp add: setsum_op_ivl_Suc [where n = "Suc n"] distrib_left lemma_realpow_diff_sumr
- del: setsum_op_ivl_Suc)
+ also have "... = (x - y) * (\<Sum>p<Suc (Suc n). x ^ p * y ^ (Suc n - p))"
+ by (simp add: field_simps Suc_diff_le setsum_left_distrib setsum_right_distrib)
finally show ?case .
-qed
+qed simp
corollary power_diff_sumr2: --{* @{text COMPLEX_POLYFUN} in HOL Light *}
fixes x :: "'a::{comm_ring,monoid_mult}"
- shows "x^n - y^n = (x - y) * (\<Sum>i=0..<n. y^(n - Suc i) * x^i)"
+ shows "x^n - y^n = (x - y) * (\<Sum>i<n. y^(n - Suc i) * x^i)"
using lemma_realpow_diff_sumr2[of x "n - 1" y]
by (cases "n = 0") (simp_all add: field_simps)
lemma lemma_realpow_rev_sumr:
- "(\<Sum>p=0..<Suc n. (x ^ p) * (y ^ (n - p))) =
- (\<Sum>p=0..<Suc n. (x ^ (n - p)) * (y ^ p))"
+ "(\<Sum>p<Suc n. (x ^ p) * (y ^ (n - p))) =
+ (\<Sum>p<Suc n. (x ^ (n - p)) * (y ^ p))"
apply (rule setsum_reindex_cong [where f="\<lambda>i. n - i"])
- apply (rule inj_onI, auto)
- apply (metis atLeastLessThan_iff diff_diff_cancel diff_less_Suc imageI le0 less_Suc_eq_le)
+ apply (auto simp: image_iff Bex_def intro!: inj_onI)
+ apply arith
done
lemma power_diff_1_eq:
fixes x :: "'a::{comm_ring,monoid_mult}"
- shows "n \<noteq> 0 \<Longrightarrow> x^n - 1 = (x - 1) * (\<Sum>i=0..<n. (x^i))"
+ shows "n \<noteq> 0 \<Longrightarrow> x^n - 1 = (x - 1) * (\<Sum>i<n. (x^i))"
using lemma_realpow_diff_sumr2 [of x _ 1]
by (cases n) auto
lemma one_diff_power_eq':
fixes x :: "'a::{comm_ring,monoid_mult}"
- shows "n \<noteq> 0 \<Longrightarrow> 1 - x^n = (1 - x) * (\<Sum>i=0..<n. x^(n - Suc i))"
+ shows "n \<noteq> 0 \<Longrightarrow> 1 - x^n = (1 - x) * (\<Sum>i<n. x^(n - Suc i))"
using lemma_realpow_diff_sumr2 [of 1 _ x]
by (cases n) auto
lemma one_diff_power_eq:
fixes x :: "'a::{comm_ring,monoid_mult}"
- shows "n \<noteq> 0 \<Longrightarrow> 1 - x^n = (1 - x) * (\<Sum>i=0..<n. x^i)"
+ shows "n \<noteq> 0 \<Longrightarrow> 1 - x^n = (1 - x) * (\<Sum>i<n. x^i)"
by (metis one_diff_power_eq' [of n x] nat_diff_setsum_reindex)
text{*Power series has a `circle` of convergence, i.e. if it sums for @{term
@@ -149,17 +138,17 @@
lemma sum_split_even_odd:
fixes f :: "nat \<Rightarrow> real"
shows
- "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) =
- (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1))"
+ "(\<Sum>i<2 * n. if even i then f i else g i) =
+ (\<Sum>i<n. f (2 * i)) + (\<Sum>i<n. g (2 * i + 1))"
proof (induct n)
case 0
then show ?case by simp
next
case (Suc n)
- have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) =
- (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1)) + (f (2 * n) + g (2 * n + 1))"
+ have "(\<Sum>i<2 * Suc n. if even i then f i else g i) =
+ (\<Sum>i<n. f (2 * i)) + (\<Sum>i<n. g (2 * i + 1)) + (f (2 * n) + g (2 * n + 1))"
using Suc.hyps unfolding One_nat_def by auto
- also have "\<dots> = (\<Sum> i = 0 ..< Suc n. f (2 * i)) + (\<Sum> i = 0 ..< Suc n. g (2 * i + 1))"
+ also have "\<dots> = (\<Sum>i<Suc n. f (2 * i)) + (\<Sum>i<Suc n. g (2 * i + 1))"
by auto
finally show ?case .
qed
@@ -173,14 +162,14 @@
fix r :: real
assume "0 < r"
from `g sums x`[unfolded sums_def, THEN LIMSEQ_D, OF this]
- obtain no where no_eq: "\<And> n. n \<ge> no \<Longrightarrow> (norm (setsum g { 0..<n } - x) < r)" by blast
-
- let ?SUM = "\<lambda> m. \<Sum> i = 0 ..< m. if even i then 0 else g ((i - 1) div 2)"
+ obtain no where no_eq: "\<And> n. n \<ge> no \<Longrightarrow> (norm (setsum g {..<n} - x) < r)" by blast
+
+ let ?SUM = "\<lambda> m. \<Sum>i<m. if even i then 0 else g ((i - 1) div 2)"
{
fix m
assume "m \<ge> 2 * no"
hence "m div 2 \<ge> no" by auto
- have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }"
+ have sum_eq: "?SUM (2 * (m div 2)) = setsum g {..< m div 2}"
using sum_split_even_odd by auto
hence "(norm (?SUM (2 * (m div 2)) - x) < r)"
using no_eq unfolding sum_eq using `m div 2 \<ge> no` by auto
@@ -220,16 +209,12 @@
have g_sums: "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x"
using sums_if'[OF `g sums x`] .
{
- have "?s 0 = 0" by auto
- have Suc_m1: "\<And> n. Suc n - 1 = n" by auto
have if_eq: "\<And>B T E. (if \<not> B then T else E) = (if B then E else T)" by auto
have "?s sums y" using sums_if'[OF `f sums y`] .
from this[unfolded sums_def, THEN LIMSEQ_Suc]
have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
- unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
- image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
- even_Suc Suc_m1 if_eq .
+ by (simp add: lessThan_Suc_eq_insert_0 image_iff setsum_reindex if_eq sums_def cong del: if_cong)
}
from sums_add[OF g_sums this] show ?thesis unfolding if_sum .
qed
@@ -239,8 +224,8 @@
lemma sums_alternating_upper_lower:
fixes a :: "nat \<Rightarrow> real"
assumes mono: "\<And>n. a (Suc n) \<le> a n" and a_pos: "\<And>n. 0 \<le> a n" and "a ----> 0"
- shows "\<exists>l. ((\<forall>n. (\<Sum>i=0..<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i=0..<2*n. -1^i*a i) ----> l) \<and>
- ((\<forall>n. l \<le> (\<Sum>i=0..<2*n + 1. -1^i*a i)) \<and> (\<lambda> n. \<Sum>i=0..<2*n + 1. -1^i*a i) ----> l)"
+ shows "\<exists>l. ((\<forall>n. (\<Sum>i<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i<2*n. -1^i*a i) ----> l) \<and>
+ ((\<forall>n. l \<le> (\<Sum>i<2*n + 1. -1^i*a i)) \<and> (\<lambda> n. \<Sum>i<2*n + 1. -1^i*a i) ----> l)"
(is "\<exists>l. ((\<forall>n. ?f n \<le> l) \<and> _) \<and> ((\<forall>n. l \<le> ?g n) \<and> _)")
proof (rule nested_sequence_unique)
have fg_diff: "\<And>n. ?f n - ?g n = - a (2 * n)" unfolding One_nat_def by auto
@@ -279,13 +264,13 @@
and a_pos: "\<And> n. 0 \<le> a n"
and a_monotone: "\<And> n. a (Suc n) \<le> a n"
shows summable: "summable (\<lambda> n. (-1)^n * a n)"
- and "\<And>n. (\<Sum>i=0..<2*n. (-1)^i*a i) \<le> (\<Sum>i. (-1)^i*a i)"
- and "(\<lambda>n. \<Sum>i=0..<2*n. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
- and "\<And>n. (\<Sum>i. (-1)^i*a i) \<le> (\<Sum>i=0..<2*n+1. (-1)^i*a i)"
- and "(\<lambda>n. \<Sum>i=0..<2*n+1. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
+ and "\<And>n. (\<Sum>i<2*n. (-1)^i*a i) \<le> (\<Sum>i. (-1)^i*a i)"
+ and "(\<lambda>n. \<Sum>i<2*n. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
+ and "\<And>n. (\<Sum>i. (-1)^i*a i) \<le> (\<Sum>i<2*n+1. (-1)^i*a i)"
+ and "(\<lambda>n. \<Sum>i<2*n+1. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
proof -
let ?S = "\<lambda>n. (-1)^n * a n"
- let ?P = "\<lambda>n. \<Sum>i=0..<n. ?S i"
+ let ?P = "\<lambda>n. \<Sum>i<n. ?S i"
let ?f = "\<lambda>n. ?P (2 * n)"
let ?g = "\<lambda>n. ?P (2 * n + 1)"
obtain l :: real
@@ -295,7 +280,7 @@
and "?g ----> l"
using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast
- let ?Sa = "\<lambda>m. \<Sum> n = 0..<m. ?S n"
+ let ?Sa = "\<lambda>m. \<Sum>n<m. ?S n"
have "?Sa ----> l"
proof (rule LIMSEQ_I)
fix r :: real
@@ -332,13 +317,13 @@
from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no"
by auto
from g[OF this] show ?thesis
- unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
+ unfolding n_eq range_eq .
qed
}
thus "\<exists>no. \<forall>n \<ge> no. norm (?Sa n - l) < r" by blast
qed
hence sums_l: "(\<lambda>i. (-1)^i * a i) sums l"
- unfolding sums_def atLeastLessThanSuc_atLeastAtMost[symmetric] .
+ unfolding sums_def .
thus "summable ?S" using summable_def by auto
have "l = suminf ?S" using sums_unique[OF sums_l] .
@@ -359,11 +344,11 @@
assumes a_zero: "a ----> 0" and "monoseq a"
shows "summable (\<lambda> n. (-1)^n * a n)" (is "?summable")
and "0 < a 0 \<longrightarrow>
- (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n. -1^i * a i .. \<Sum>i=0..<2*n+1. -1^i * a i})" (is "?pos")
+ (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i<2*n. -1^i * a i .. \<Sum>i<2*n+1. -1^i * a i})" (is "?pos")
and "a 0 < 0 \<longrightarrow>
- (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n+1. -1^i * a i .. \<Sum>i=0..<2*n. -1^i * a i})" (is "?neg")
- and "(\<lambda>n. \<Sum>i=0..<2*n. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?f")
- and "(\<lambda>n. \<Sum>i=0..<2*n+1. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?g")
+ (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i<2*n+1. -1^i * a i .. \<Sum>i<2*n. -1^i * a i})" (is "?neg")
+ and "(\<lambda>n. \<Sum>i<2*n. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?f")
+ and "(\<lambda>n. \<Sum>i<2*n+1. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?g")
proof -
have "?summable \<and> ?pos \<and> ?neg \<and> ?f \<and> ?g"
proof (cases "(\<forall> n. 0 \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m)")
@@ -424,38 +409,32 @@
subsection {* Term-by-Term Differentiability of Power Series *}
-definition diffs :: "(nat => 'a::ring_1) => nat => 'a"
- where "diffs c = (\<lambda>n. of_nat (Suc n) * c(Suc n))"
+definition diffs :: "(nat \<Rightarrow> 'a::ring_1) \<Rightarrow> nat \<Rightarrow> 'a"
+ where "diffs c = (\<lambda>n. of_nat (Suc n) * c (Suc n))"
text{*Lemma about distributing negation over it*}
lemma diffs_minus: "diffs (\<lambda>n. - c n) = (\<lambda>n. - diffs c n)"
by (simp add: diffs_def)
lemma sums_Suc_imp:
- assumes f: "f 0 = 0"
- shows "(\<lambda>n. f (Suc n)) sums s \<Longrightarrow> (\<lambda>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
+ "(f::nat \<Rightarrow> 'a::real_normed_vector) 0 = 0 \<Longrightarrow> (\<lambda>n. f (Suc n)) sums s \<Longrightarrow> (\<lambda>n. f n) sums s"
+ using sums_Suc_iff[of f] by simp
lemma diffs_equiv:
fixes x :: "'a::{real_normed_vector, ring_1}"
- shows "summable (\<lambda>n. (diffs c)(n) * (x ^ n)) \<Longrightarrow>
- (\<lambda>n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums
- (\<Sum>n. (diffs c)(n) * (x ^ n))"
+ shows "summable (\<lambda>n. diffs c n * x^n) \<Longrightarrow>
+ (\<lambda>n. of_nat n * c n * x^(n - Suc 0)) sums (\<Sum>n. diffs c n * x^n)"
unfolding diffs_def
by (simp add: summable_sums sums_Suc_imp)
lemma lemma_termdiff1:
fixes z :: "'a :: {monoid_mult,comm_ring}" shows
- "(\<Sum>p=0..<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =
- (\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
+ "(\<Sum>p<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =
+ (\<Sum>p<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
by (auto simp add: algebra_simps power_add [symmetric])
lemma sumr_diff_mult_const2:
- "setsum f {0..<n} - of_nat n * (r::'a::ring_1) = (\<Sum>i = 0..<n. f i - r)"
+ "setsum f {..<n} - of_nat n * (r::'a::ring_1) = (\<Sum>i<n. f i - r)"
by (simp add: setsum_subtractf)
lemma lemma_termdiff2:
@@ -463,7 +442,7 @@
assumes h: "h \<noteq> 0"
shows
"((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) =
- h * (\<Sum>p=0..< n - Suc 0. \<Sum>q=0..< n - Suc 0 - p.
+ h * (\<Sum>p< n - Suc 0. \<Sum>q< n - Suc 0 - p.
(z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs")
apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h)
apply (simp add: right_diff_distrib diff_divide_distrib h)
@@ -471,7 +450,7 @@
apply (cases "n", simp)
apply (simp add: lemma_realpow_diff_sumr2 h
right_diff_distrib [symmetric] mult_assoc
- del: power_Suc setsum_op_ivl_Suc of_nat_Suc)
+ del: power_Suc setsum_lessThan_Suc of_nat_Suc)
apply (subst lemma_realpow_rev_sumr)
apply (subst sumr_diff_mult_const2)
apply simp
@@ -480,7 +459,7 @@
apply (simp add: less_iff_Suc_add)
apply (clarify)
apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
- del: setsum_op_ivl_Suc power_Suc)
+ del: setsum_lessThan_Suc power_Suc)
apply (subst mult_assoc [symmetric], subst power_add [symmetric])
apply (simp add: mult_ac)
done
@@ -489,7 +468,7 @@
fixes K :: "'a::linordered_semidom"
assumes f: "\<And>p::nat. p < n \<Longrightarrow> f p \<le> K"
and K: "0 \<le> K"
- shows "setsum f {0..<n-k} \<le> of_nat n * K"
+ shows "setsum f {..<n-k} \<le> of_nat n * K"
apply (rule order_trans [OF setsum_mono])
apply (rule f, simp)
apply (simp add: mult_right_mono K)
@@ -504,7 +483,7 @@
\<le> of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h"
proof -
have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) =
- norm (\<Sum>p = 0..<n - Suc 0. \<Sum>q = 0..<n - Suc 0 - p.
+ norm (\<Sum>p<n - Suc 0. \<Sum>q<n - Suc 0 - p.
(z + h) ^ q * z ^ (n - 2 - q)) * norm h"
by (metis (lifting, no_types) lemma_termdiff2 [OF 1] mult_commute norm_mult)
also have "\<dots> \<le> of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h"
@@ -516,7 +495,7 @@
apply (simp only: norm_mult norm_power power_add)
apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K)
done
- show "norm (\<Sum>p = 0..<n - Suc 0. \<Sum>q = 0..<n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))
+ show "norm (\<Sum>p<n - Suc 0. \<Sum>q<n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))
\<le> of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))"
apply (intro
order_trans [OF norm_setsum]
@@ -712,16 +691,16 @@
from divide_pos_pos[OF `0 < r` this]
have "0 < ?r" .
- let "?s n" = "SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
- def S' \<equiv> "Min (?s ` { 0 ..< ?N })"
+ let ?s = "\<lambda>n. SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
+ def S' \<equiv> "Min (?s ` {..< ?N })"
have "0 < S'" unfolding S'_def
proof (rule iffD2[OF Min_gr_iff])
- show "\<forall>x \<in> (?s ` { 0 ..< ?N }). 0 < x"
+ show "\<forall>x \<in> (?s ` {..< ?N }). 0 < x"
proof
fix x
- assume "x \<in> ?s ` {0..<?N}"
- then obtain n where "x = ?s n" and "n \<in> {0..<?N}"
+ assume "x \<in> ?s ` {..<?N}"
+ then obtain n where "x = ?s n" and "n \<in> {..<?N}"
using image_iff[THEN iffD1] by blast
from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def]
obtain s where s_bound: "0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < s \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r)"
@@ -750,32 +729,30 @@
note div_shft_smbl = summable_divide[OF diff_shft_smbl]
note all_shft_smbl = summable_diff[OF div_smbl ign[OF `summable (f' x0)`]]
- {
- fix n
+ { fix n
have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>"
using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero]
unfolding abs_divide .
hence "\<bar> (\<bar>?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)"
- using `x \<noteq> 0` by auto
- } note L_ge = summable_le2[OF allI[OF this] ign[OF `summable L`]]
- from order_trans[OF summable_rabs[OF conjunct1[OF L_ge]] L_ge[THEN conjunct2]]
- have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> (\<Sum> i. L (i + ?N))" .
- hence "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> r / 3" (is "?L_part \<le> r/3")
+ using `x \<noteq> 0` by auto }
+ note 1 = this and 2 = summable_rabs_comparison_test[OF _ ign[OF `summable L`]]
+ then have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> (\<Sum> i. L (i + ?N))"
+ by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] summable_le[OF _ 2 ign[OF `summable L`]]])
+ then have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> r / 3" (is "?L_part \<le> r/3")
using L_estimate by auto
- have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> \<le>
- (\<Sum>n \<in> { 0 ..< ?N}. \<bar>?diff n x - f' x0 n \<bar>)" ..
- also have "\<dots> < (\<Sum>n \<in> { 0 ..< ?N}. ?r)"
+ have "\<bar>\<Sum>n<?N. ?diff n x - f' x0 n \<bar> \<le> (\<Sum>n<?N. \<bar>?diff n x - f' x0 n \<bar>)" ..
+ also have "\<dots> < (\<Sum>n<?N. ?r)"
proof (rule setsum_strict_mono)
fix n
- assume "n \<in> { 0 ..< ?N}"
+ assume "n \<in> {..< ?N}"
have "\<bar>x\<bar> < S" using `\<bar>x\<bar> < S` .
also have "S \<le> S'" using `S \<le> S'` .
also have "S' \<le> ?s n" unfolding S'_def
proof (rule Min_le_iff[THEN iffD2])
- have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n"
- using `n \<in> { 0 ..< ?N}` by auto
- thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
+ have "?s n \<in> (?s ` {..<?N}) \<and> ?s n \<le> ?s n"
+ using `n \<in> {..< ?N}` by auto
+ thus "\<exists> a \<in> (?s ` {..<?N}). a \<le> ?s n" by blast
qed auto
finally have "\<bar>x\<bar> < ?s n" .
@@ -784,12 +761,12 @@
with `x \<noteq> 0` and `\<bar>x\<bar> < ?s n` show "\<bar>?diff n x - f' x0 n\<bar> < ?r"
by blast
qed auto
- also have "\<dots> = of_nat (card {0 ..< ?N}) * ?r"
+ also have "\<dots> = of_nat (card {..<?N}) * ?r"
by (rule setsum_constant)
also have "\<dots> = real ?N * ?r"
unfolding real_eq_of_nat by auto
also have "\<dots> = r/3" by auto
- finally have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> < r / 3" (is "?diff_part < r / 3") .
+ finally have "\<bar>\<Sum>n<?N. ?diff n x - f' x0 n \<bar> < r / 3" (is "?diff_part < r / 3") .
from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
have "\<bar>(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\<bar> =
@@ -799,6 +776,7 @@
also have "\<dots> \<le> ?diff_part + \<bar> (\<Sum>n. ?diff (n + ?N) x) - (\<Sum> n. f' x0 (n + ?N)) \<bar>"
unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"]
unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]]
+ apply (subst (5) add_commute)
by (rule abs_triangle_ineq)
also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part"
using abs_triangle_ineq4 by auto
@@ -844,17 +822,17 @@
show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
proof -
have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> =
- (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>"
+ (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p<Suc n. x ^ p * y ^ (n - p)\<bar>"
unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult
by auto
also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)"
proof (rule mult_left_mono)
- have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)"
+ have "\<bar>\<Sum>p<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)"
by (rule setsum_abs)
- also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
+ also have "\<dots> \<le> (\<Sum>p<Suc n. R' ^ n)"
proof (rule setsum_mono)
fix p
- assume "p \<in> {0..<Suc n}"
+ assume "p \<in> {..<Suc n}"
hence "p \<le> n" by auto
{
fix n
@@ -872,7 +850,7 @@
qed
also have "\<dots> = real (Suc n) * R' ^ n"
unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
- finally show "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>"
+ finally show "\<bar>\<Sum>p<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>"
unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>"
unfolding abs_mult[symmetric] by auto
@@ -893,14 +871,14 @@
hence "R' \<in> {-R <..< R}" and "norm x < norm R'"
using assms `R' < R` by auto
have "summable (\<lambda> n. f n * x^n)"
- proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
+ proof (rule summable_comparison_test, intro exI allI impI)
fix n
have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)"
by (rule mult_left_mono) auto
- show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)"
+ show "norm (f n * x ^ n) \<le> norm (f n * real (Suc n) * x ^ n)"
unfolding real_norm_def abs_mult
by (rule mult_right_mono) (auto simp add: le[unfolded mult_1_right])
- qed
+ qed (rule powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`])
from this[THEN summable_mult2[where c=x], unfolded mult_assoc, unfolded mult_commute]
show "summable (?f x)" by auto
}
@@ -947,7 +925,7 @@
obtain N :: nat where N: "norm x < real N * r"
using reals_Archimedean3 [OF r0] by fast
from r1 show ?thesis
- proof (rule ratio_test [rule_format])
+ proof (rule summable_ratio_test [rule_format])
fix n :: nat
assume n: "N \<le> n"
have "norm x \<le> real N * r"
@@ -1027,7 +1005,7 @@
fixes f :: "nat \<Rightarrow> 'a::{real_normed_algebra_1}"
shows "(\<Sum>n. f n * 0 ^ n) = f 0"
proof -
- have "(\<Sum>n = 0..<1. f n * 0 ^ n) = (\<Sum>n. f n * 0 ^ n)"
+ have "(\<Sum>n<1. f n * 0 ^ n) = (\<Sum>n. f n * 0 ^ n)"
by (rule sums_unique [OF series_zero], simp add: power_0_left)
thus ?thesis unfolding One_nat_def by simp
qed
@@ -1147,7 +1125,7 @@
using order_le_imp_less_or_eq [OF assms]
proof
assume "0 < x"
- have "1+x \<le> (\<Sum>n = 0..<2. inverse (real (fact n)) * x ^ n)"
+ have "1+x \<le> (\<Sum>n<2. inverse (real (fact n)) * x ^ n)"
by (auto simp add: numeral_2_eq_2)
also have "... \<le> (\<Sum>n. inverse (real (fact n)) * x ^ n)"
apply (rule series_pos_le [OF summable_exp])
@@ -1395,12 +1373,13 @@
proof -
have "exp x = suminf (\<lambda>n. inverse(fact n) * (x ^ n))"
by (simp add: exp_def)
- also from summable_exp have "... = (\<Sum> n::nat = 0 ..< 2. inverse(fact n) * (x ^ n)) +
- (\<Sum> n. inverse(fact(n+2)) * (x ^ (n+2)))" (is "_ = ?a + _")
+ also from summable_exp have "... = (\<Sum> n. inverse(fact(n+2)) * (x ^ (n+2))) +
+ (\<Sum> n::nat<2. inverse(fact n) * (x ^ n))" (is "_ = _ + ?a")
by (rule suminf_split_initial_segment)
also have "?a = 1 + x"
by (simp add: numeral_2_eq_2)
- finally show ?thesis .
+ finally show ?thesis
+ by simp
qed
lemma exp_bound: "0 <= (x::real) \<Longrightarrow> x <= 1 \<Longrightarrow> exp x <= 1 + x + x\<^sup>2"
@@ -2433,6 +2412,25 @@
lemmas realpow_num_eq_if = power_eq_if
+lemma sumr_pos_lt_pair:
+ fixes f :: "nat \<Rightarrow> real"
+ shows "\<lbrakk>summable f;
+ \<And>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc(Suc 0) * d) + 1))\<rbrakk>
+ \<Longrightarrow> setsum f {..<k} < suminf f"
+unfolding One_nat_def
+apply (subst suminf_split_initial_segment [where k="k"])
+apply assumption
+apply simp
+apply (drule_tac k="k" in summable_ignore_initial_segment)
+apply (drule_tac k="Suc (Suc 0)" in sums_group [OF summable_sums], simp)
+apply simp
+apply (frule sums_unique)
+apply (drule sums_summable)
+apply simp
+apply (erule suminf_gt_zero)
+apply (simp add: add_ac)
+done
+
lemma cos_two_less_zero [simp]:
"cos 2 < 0"
proof -
@@ -2444,9 +2442,9 @@
by simp
then have **: "summable (\<lambda>n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
by (rule sums_summable)
- have "0 < (\<Sum>n = 0..<Suc (Suc (Suc 0)). - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
+ have "0 < (\<Sum>n<Suc (Suc (Suc 0)). - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
by (simp add: fact_num_eq_if_nat realpow_num_eq_if)
- moreover have "(\<Sum>n = 0..<Suc (Suc (Suc 0)). - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))
+ moreover have "(\<Sum>n<Suc (Suc (Suc 0)). - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))
< (\<Sum>n. - (-1 ^ n * 2 ^ (2 * n) / real (fact (2 * n))))"
proof -
{ fix d
@@ -3807,7 +3805,7 @@
case False
hence "\<bar>x\<bar> = 1" using `\<bar>x\<bar> \<le> 1` by auto
let ?a = "\<lambda>x n. \<bar>1 / real (n*2+1) * x^(n*2+1)\<bar>"
- let ?diff = "\<lambda> x n. \<bar> arctan x - (\<Sum> i = 0..<n. ?c x i)\<bar>"
+ let ?diff = "\<lambda> x n. \<bar> arctan x - (\<Sum> i<n. ?c x i)\<bar>"
{
fix n :: nat
have "0 < (1 :: real)" by auto
@@ -3830,7 +3828,7 @@
from `even n` obtain m where "2 * m = n"
unfolding even_mult_two_ex by auto
from bounds[of m, unfolded this atLeastAtMost_iff]
- have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n + 1. (?c x i)) - (\<Sum>i = 0..<n. (?c x i))"
+ have "\<bar>arctan x - (\<Sum>i<n. (?c x i))\<bar> \<le> (\<Sum>i<n + 1. (?c x i)) - (\<Sum>i<n. (?c x i))"
by auto
also have "\<dots> = ?c x n" unfolding One_nat_def by auto
also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
@@ -3842,7 +3840,7 @@
unfolding odd_Suc_mult_two_ex by auto
hence m_plus: "2 * (m + 1) = n + 1" by auto
from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
- have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n. (?c x i)) - (\<Sum>i = 0..<n+1. (?c x i))"
+ have "\<bar>arctan x - (\<Sum>i<n. (?c x i))\<bar> \<le> (\<Sum>i<n. (?c x i)) - (\<Sum>i<n+1. (?c x i))"
by auto
also have "\<dots> = - ?c x n" unfolding One_nat_def by auto
also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto