--- a/src/Doc/Logics/document/HOL.tex Mon Oct 17 11:07:01 2016 +0200
+++ b/src/Doc/Logics/document/HOL.tex Mon Oct 17 14:37:32 2016 +0200
@@ -145,7 +145,7 @@
numbers). The type class \cldx{plus_ac0} comprises all types for which 0
and~+ satisfy the laws $x+y=y+x$, $(x+y)+z = x+(y+z)$ and $0+x = x$. These
types include the numeric ones \tdx{nat}, \tdx{int} and~\tdx{real} and also
-multisets. The summation operator \cdx{setsum} is available for all types in
+multisets. The summation operator \cdx{sum} is available for all types in
this class.
Theory \thydx{Ord} defines the syntactic class \cldx{ord} of order
--- a/src/Doc/Main/Main_Doc.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/Doc/Main/Main_Doc.thy Mon Oct 17 14:37:32 2016 +0200
@@ -408,7 +408,7 @@
@{const Finite_Set.finite} & @{term_type_only Finite_Set.finite "'a set\<Rightarrow>bool"}\\
@{const Finite_Set.card} & @{term_type_only Finite_Set.card "'a set \<Rightarrow> nat"}\\
@{const Finite_Set.fold} & @{term_type_only Finite_Set.fold "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b"}\\
-@{const Groups_Big.setsum} & @{term_type_only Groups_Big.setsum "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b::comm_monoid_add"}\\
+@{const Groups_Big.sum} & @{term_type_only Groups_Big.sum "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b::comm_monoid_add"}\\
@{const Groups_Big.setprod} & @{term_type_only Groups_Big.setprod "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b::comm_monoid_mult"}\\
\end{supertabular}
@@ -416,8 +416,8 @@
\subsubsection*{Syntax}
\begin{supertabular}{@ {} l @ {\quad$\equiv$\quad} l l @ {}}
-@{term "setsum (\<lambda>x. x) A"} & @{term[source]"setsum (\<lambda>x. x) A"} & (\<^verbatim>\<open>SUM\<close>)\\
-@{term "setsum (\<lambda>x. t) A"} & @{term[source]"setsum (\<lambda>x. t) A"}\\
+@{term "sum (\<lambda>x. x) A"} & @{term[source]"sum (\<lambda>x. x) A"} & (\<^verbatim>\<open>SUM\<close>)\\
+@{term "sum (\<lambda>x. t) A"} & @{term[source]"sum (\<lambda>x. t) A"}\\
@{term[source] "\<Sum>x|P. t"} & @{term"\<Sum>x|P. t"}\\
\multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Prod>\<close> instead of \<open>\<Sum>\<close>} & (\<^verbatim>\<open>PROD\<close>)\\
\end{supertabular}
@@ -463,10 +463,10 @@
@{term[source] "\<Union>i\<le>n. A"} & @{term[source] "\<Union>i \<in> {..n}. A"}\\
@{term[source] "\<Union>i<n. A"} & @{term[source] "\<Union>i \<in> {..<n}. A"}\\
\multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Inter>\<close> instead of \<open>\<Union>\<close>}\\
-@{term "setsum (\<lambda>x. t) {a..b}"} & @{term[source] "setsum (\<lambda>x. t) {a..b}"}\\
-@{term "setsum (\<lambda>x. t) {a..<b}"} & @{term[source] "setsum (\<lambda>x. t) {a..<b}"}\\
-@{term "setsum (\<lambda>x. t) {..b}"} & @{term[source] "setsum (\<lambda>x. t) {..b}"}\\
-@{term "setsum (\<lambda>x. t) {..<b}"} & @{term[source] "setsum (\<lambda>x. t) {..<b}"}\\
+@{term "sum (\<lambda>x. t) {a..b}"} & @{term[source] "sum (\<lambda>x. t) {a..b}"}\\
+@{term "sum (\<lambda>x. t) {a..<b}"} & @{term[source] "sum (\<lambda>x. t) {a..<b}"}\\
+@{term "sum (\<lambda>x. t) {..b}"} & @{term[source] "sum (\<lambda>x. t) {..b}"}\\
+@{term "sum (\<lambda>x. t) {..<b}"} & @{term[source] "sum (\<lambda>x. t) {..<b}"}\\
\multicolumn{2}{@ {}l@ {}}{Similarly for \<open>\<Prod>\<close> instead of \<open>\<Sum>\<close>}\\
\end{supertabular}
--- a/src/Doc/Prog_Prove/Bool_nat_list.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/Doc/Prog_Prove/Bool_nat_list.thy Mon Oct 17 14:37:32 2016 +0200
@@ -453,9 +453,9 @@
\end{exercise}
\begin{exercise}
-Define a recursive function @{text "sum ::"} @{typ"nat \<Rightarrow> nat"} such that
-\mbox{@{text"sum n"}} @{text"="} @{text"0 + ... + n"} and prove
-@{prop" sum(n::nat) = n * (n+1) div 2"}.
+Define a recursive function @{text "sum_upto ::"} @{typ"nat \<Rightarrow> nat"} such that
+\mbox{@{text"sum_upto n"}} @{text"="} @{text"0 + ... + n"} and prove
+@{prop" sum_upto (n::nat) = n * (n+1) div 2"}.
\end{exercise}
*}
(*<*)
--- a/src/Doc/Sugar/Sugar.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/Doc/Sugar/Sugar.thy Mon Oct 17 14:37:32 2016 +0200
@@ -13,9 +13,9 @@
This document is for those Isabelle users who have mastered
the art of mixing \LaTeX\ text and Isabelle theories and never want to
typeset a theorem by hand anymore because they have experienced the
-bliss of writing \verb!@!\verb!{thm[display,mode=latex_sum] setsum_Suc_diff [no_vars]}!
+bliss of writing \verb!@!\verb!{thm[display,mode=latex_sum] sum_Suc_diff [no_vars]}!
and seeing Isabelle typeset it for them:
-@{thm[display,mode=latex_sum] setsum_Suc_diff[no_vars]}
+@{thm[display,mode=latex_sum] sum_Suc_diff[no_vars]}
No typos, no omissions, no sweat.
If you have not experienced that joy, read Chapter 4, \emph{Presenting
Theories}, @{cite LNCS2283} first.
--- a/src/HOL/Algebra/IntRing.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Algebra/IntRing.thy Mon Oct 17 14:37:32 2016 +0200
@@ -91,7 +91,7 @@
rewrites int_carrier_eq: "carrier \<Z> = UNIV"
and int_zero_eq: "zero \<Z> = 0"
and int_add_eq: "add \<Z> x y = x + y"
- and int_finsum_eq: "finsum \<Z> f A = setsum f A"
+ and int_finsum_eq: "finsum \<Z> f A = sum f A"
proof -
\<comment> "Specification"
show "abelian_monoid \<Z>" by standard auto
@@ -105,7 +105,7 @@
note add = this
show zero: "zero \<Z> = 0"
by simp
- show "finsum \<Z> f A = setsum f A"
+ show "finsum \<Z> f A = sum f A"
by (induct A rule: infinite_finite_induct, auto)
qed
@@ -117,7 +117,7 @@
rewrites "carrier \<Z> = UNIV"
and "zero \<Z> = 0"
and "add \<Z> x y = x + y"
- and "finsum \<Z> f A = setsum f A"
+ and "finsum \<Z> f A = sum f A"
and int_a_inv_eq: "a_inv \<Z> x = - x"
and int_a_minus_eq: "a_minus \<Z> x y = x - y"
proof -
@@ -150,7 +150,7 @@
rewrites "carrier \<Z> = UNIV"
and "zero \<Z> = 0"
and "add \<Z> x y = x + y"
- and "finsum \<Z> f A = setsum f A"
+ and "finsum \<Z> f A = sum f A"
and "a_inv \<Z> x = - x"
and "a_minus \<Z> x y = x - y"
proof -
--- a/src/HOL/Analysis/Binary_Product_Measure.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Binary_Product_Measure.thy Mon Oct 17 14:37:32 2016 +0200
@@ -767,8 +767,8 @@
apply (subst B.emeasure_pair_measure_alt[OF X])
apply (subst emeasure_count_space)
apply (auto simp add: emeasure_count_space nn_integral_count_space
- pos_card of_nat_setsum[symmetric] card_SigmaI[symmetric]
- simp del: of_nat_setsum card_SigmaI
+ pos_card of_nat_sum[symmetric] card_SigmaI[symmetric]
+ simp del: of_nat_sum card_SigmaI
intro!: arg_cong[where f=card])
done
qed
--- a/src/HOL/Analysis/Bochner_Integration.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Bochner_Integration.thy Mon Oct 17 14:37:32 2016 +0200
@@ -157,10 +157,10 @@
lemma
fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
- shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
- and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
+ shows sum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
+ and sum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
unfolding indicator_def
- using assms by (auto intro!: setsum.mono_neutral_cong_right split: if_split_asm)
+ using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm)
lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool"
@@ -220,8 +220,8 @@
next
case (insert x F)
then show ?case
- by (auto intro!: add mult set setsum_nonneg split: split_indicator split_indicator_asm
- simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff)
+ by (auto intro!: add mult set sum_nonneg split: split_indicator split_indicator_asm
+ simp del: sum_mult_indicator simp: sum_nonneg_eq_0_iff)
qed
with U' show "P (U' i)" by simp
qed
@@ -337,7 +337,7 @@
(\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} else 0) *\<^sub>R y)"
unfolding simple_bochner_integral_def
- proof (safe intro!: setsum.cong scaleR_cong_right)
+ proof (safe intro!: sum.cong scaleR_cong_right)
fix y assume y: "y \<in> space M" "f y \<noteq> 0"
have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
@@ -359,17 +359,17 @@
ultimately
show "measure M {x \<in> space M. f x = f y} =
(\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)"
- apply (simp add: setsum.If_cases eq)
+ apply (simp add: sum.If_cases eq)
apply (subst measure_finite_Union[symmetric])
apply (auto simp: disjoint_family_on_def less_top)
done
qed
also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))"
- by (auto intro!: setsum.cong simp: scaleR_setsum_left)
+ by (auto intro!: sum.cong simp: scaleR_sum_left)
also have "\<dots> = ?r"
- by (subst setsum.commute)
- (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
+ by (subst sum.commute)
+ (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
finally show "simple_bochner_integral M f = ?r" .
qed
@@ -391,7 +391,7 @@
by (intro simple_bochner_integral_partition)
(auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
ultimately show ?thesis
- by (simp add: setsum.distrib[symmetric] scaleR_add_right)
+ by (simp add: sum.distrib[symmetric] scaleR_add_right)
qed
lemma (in linear) simple_bochner_integral_linear:
@@ -404,7 +404,7 @@
(auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"] zero
elim: simple_bochner_integrable.cases)
also have "\<dots> = f (simple_bochner_integral M g)"
- by (simp add: simple_bochner_integral_def setsum scaleR)
+ by (simp add: simple_bochner_integral_def sum scaleR)
finally show ?thesis .
qed
@@ -431,7 +431,7 @@
proof -
have "norm (simple_bochner_integral M f) \<le>
(\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))"
- unfolding simple_bochner_integral_def by (rule norm_setsum)
+ unfolding simple_bochner_integral_def by (rule norm_sum)
also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)"
by simp
also have "\<dots> = simple_bochner_integral M (\<lambda>x. norm (f x))"
@@ -444,7 +444,7 @@
lemma simple_bochner_integral_nonneg[simp]:
fixes f :: "'a \<Rightarrow> real"
shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f"
- by (simp add: setsum_nonneg simple_bochner_integral_def)
+ by (simp add: sum_nonneg simple_bochner_integral_def)
lemma simple_bochner_integral_eq_nn_integral:
assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x"
@@ -473,9 +473,9 @@
unfolding simple_integral_def
by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
(auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
- intro!: setsum.cong ennreal_cong_mult
- simp: setsum_ennreal[symmetric] ac_simps ennreal_mult
- simp del: setsum_ennreal)
+ intro!: sum.cong ennreal_cong_mult
+ simp: sum_ennreal[symmetric] ac_simps ennreal_mult
+ simp del: sum_ennreal)
also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
using f
by (intro nn_integral_eq_simple_integral[symmetric])
@@ -714,7 +714,7 @@
unfolding diff_conv_add_uminus
by (intro has_bochner_integral_add has_bochner_integral_minus)
-lemma has_bochner_integral_setsum:
+lemma has_bochner_integral_sum:
"(\<And>i. i \<in> I \<Longrightarrow> has_bochner_integral M (f i) (x i)) \<Longrightarrow>
has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)"
by (induct I rule: infinite_finite_induct) auto
@@ -882,7 +882,7 @@
by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2
simple_bochner_integral_def Collect_restrict
split: split_indicator split_indicator_asm
- intro!: setsum.mono_neutral_cong_left arg_cong2[where f=measure])
+ intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure])
qed
context
@@ -971,8 +971,8 @@
lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
by (metis has_bochner_integral_zero integrable.simps)
-lemma integrable_setsum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
- by (metis has_bochner_integral_setsum integrable.simps)
+lemma integrable_sum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
+ by (metis has_bochner_integral_sum integrable.simps)
lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
integrable M (\<lambda>x. indicator A x *\<^sub>R c)"
@@ -1045,13 +1045,13 @@
integral\<^sup>L M (\<lambda>x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g"
by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable)
-lemma integral_setsum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
+lemma integral_sum: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
- by (intro has_bochner_integral_integral_eq has_bochner_integral_setsum has_bochner_integral_integrable)
-
-lemma integral_setsum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
+ by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable)
+
+lemma integral_sum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
- unfolding simp_implies_def by (rule integral_setsum)
+ unfolding simp_implies_def by (rule integral_sum)
lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
@@ -1729,9 +1729,9 @@
using summable
proof eventually_elim
fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))"
- have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_setsum)
+ have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_sum)
also have "\<dots> \<le> (\<Sum>i. norm (f i x))"
- using setsum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto
+ using sum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto
finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp
qed
@@ -1835,7 +1835,7 @@
have "(\<And>i::'b. i \<in> A \<Longrightarrow> integrable M (f i::'a \<Rightarrow> 'b)) \<Longrightarrow>
(\<And>i. i \<in> A \<Longrightarrow> P (f i)) \<Longrightarrow> P (\<lambda>x. \<Sum>i\<in>A. f i x)"
by (induct A rule: infinite_finite_induct) (auto intro!: add) }
- note setsum = this
+ note sum = this
define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z
then have s'_eq_s: "\<And>i x. x \<in> space M \<Longrightarrow> s' i x = s i x"
@@ -1873,7 +1873,7 @@
using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top)
finally have "emeasure M {y \<in> space M. s' i y = s' i x} \<noteq> \<infinity>" by simp }
then show "P (s' i)"
- by (subst s'_eq) (auto intro!: setsum base simp: less_top)
+ by (subst s'_eq) (auto intro!: sum base simp: less_top)
fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x"
by (simp add: s'_eq_s)
@@ -2191,19 +2191,19 @@
shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"
proof -
have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)"
- by (intro setsum.mono_neutral_cong_left) auto
+ by (intro sum.mono_neutral_cong_left) auto
have "(\<integral>x. f x \<partial>count_space A) = (\<integral>x. (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. indicator {a} x *\<^sub>R f a) \<partial>count_space A)"
by (intro integral_cong refl) (simp add: f eq)
also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)"
- by (subst integral_setsum) (auto intro!: setsum.cong)
+ by (subst integral_sum) (auto intro!: sum.cong)
finally show ?thesis
by auto
qed
lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
by (subst lebesgue_integral_count_space_finite_support)
- (auto intro!: setsum.mono_neutral_cong_left)
+ (auto intro!: sum.mono_neutral_cong_left)
lemma integrable_count_space_nat_iff:
fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
@@ -2491,7 +2491,7 @@
by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong)
qed
also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})"
- using finite by (subst integral_setsum) (auto simp add: integrable_mult_left_iff)
+ using finite by (subst integral_sum) (auto simp add: integrable_mult_left_iff)
finally show ?thesis .
qed
@@ -2645,7 +2645,7 @@
(\<Sum>z\<in>s i ` (space N \<times> space M). measure M {y \<in> space M. s i (x, y) = z} *\<^sub>R z)"
using s(1)[THEN simple_functionD(1)]
unfolding simple_bochner_integral_def
- by (intro setsum.mono_neutral_cong_left)
+ by (intro sum.mono_neutral_cong_left)
(auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) }
note eq = this
--- a/src/HOL/Analysis/Borel_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Borel_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -969,7 +969,7 @@
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
-lemma borel_measurable_setsum[measurable (raw)]:
+lemma borel_measurable_sum[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
@@ -1646,7 +1646,7 @@
and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
using assms by (simp_all add: minus_ereal_def divide_ereal_def)
-lemma borel_measurable_ereal_setsum[measurable (raw)]:
+lemma borel_measurable_ereal_sum[measurable (raw)]:
fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
--- a/src/HOL/Analysis/Bounded_Continuous_Function.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Bounded_Continuous_Function.thy Mon Oct 17 14:37:32 2016 +0200
@@ -426,7 +426,7 @@
by (simp add: cbox_def)
then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le>
(\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)"
- by (auto intro!: setsum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric])
+ by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric])
then show ?thesis
by (auto intro: real_sqrt_le_mono
simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] setL2_def)
--- a/src/HOL/Analysis/Bounded_Linear_Function.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Bounded_Linear_Function.thy Mon Oct 17 14:37:32 2016 +0200
@@ -43,7 +43,7 @@
bounded_linear_scaleR_const
bounded_linear_const_scaleR
bounded_linear_ident
- bounded_linear_setsum
+ bounded_linear_sum
bounded_linear_Pair
bounded_linear_sub
bounded_linear_fst_comp
@@ -204,8 +204,8 @@
scaleR_right
zero_left
zero_right
- setsum_left
- setsum_right
+ sum_left
+ sum_right
end
@@ -340,20 +340,20 @@
subsection \<open>On Euclidean Space\<close>
-lemma Zfun_setsum:
+lemma Zfun_sum:
assumes "finite s"
assumes f: "\<And>i. i \<in> s \<Longrightarrow> Zfun (f i) F"
- shows "Zfun (\<lambda>x. setsum (\<lambda>i. f i x) s) F"
+ shows "Zfun (\<lambda>x. sum (\<lambda>i. f i x) s) F"
using assms by induct (auto intro!: Zfun_zero Zfun_add)
lemma norm_blinfun_euclidean_le:
fixes a::"'a::euclidean_space \<Rightarrow>\<^sub>L 'b::real_normed_vector"
- shows "norm a \<le> setsum (\<lambda>x. norm (a x)) Basis"
+ shows "norm a \<le> sum (\<lambda>x. norm (a x)) Basis"
apply (rule norm_blinfun_bound)
- apply (simp add: setsum_nonneg)
+ apply (simp add: sum_nonneg)
apply (subst euclidean_representation[symmetric, where 'a='a])
- apply (simp only: blinfun.bilinear_simps setsum_distrib_right)
- apply (rule order.trans[OF norm_setsum setsum_mono])
+ apply (simp only: blinfun.bilinear_simps sum_distrib_right)
+ apply (rule order.trans[OF norm_sum sum_mono])
apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)
done
@@ -366,7 +366,7 @@
have "\<And>j. j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x j - a j)) F"
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
hence "Zfun (\<lambda>x. \<Sum>j\<in>Basis. norm (b x j - a j)) F"
- by (auto intro!: Zfun_setsum)
+ by (auto intro!: Zfun_sum)
thus ?thesis
unfolding tendsto_Zfun_iff
by (rule Zfun_le)
@@ -387,9 +387,9 @@
have "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b
= (\<Sum>j\<in>Basis. if j = b then (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j))) else 0)"
using b
- by (auto simp add: algebra_simps inner_setsum_left inner_Basis split: if_split intro!: setsum.cong)
+ by (auto simp add: algebra_simps inner_sum_left inner_Basis split: if_split intro!: sum.cong)
also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> b)))"
- using b by (simp add: setsum.delta)
+ using b by (simp add: sum.delta)
also have "\<dots> = f x \<bullet> b"
by (metis (mono_tags, lifting) Linear_Algebra.linear_componentwise linear_axioms)
finally show "(\<Sum>j\<in>Basis. \<Sum>i\<in>Basis. (x \<bullet> i * (f i \<bullet> j)) *\<^sub>R j) \<bullet> b = f x \<bullet> b" .
@@ -400,15 +400,15 @@
by transfer simp
lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)"
- by transfer (auto simp: algebra_simps setsum_subtractf)
+ by transfer (auto simp: algebra_simps sum_subtractf)
lemma norm_blinfun_of_matrix:
"norm (blinfun_of_matrix a) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>a i j\<bar>)"
apply (rule norm_blinfun_bound)
- apply (simp add: setsum_nonneg)
- apply (simp only: blinfun_of_matrix_apply setsum_distrib_right)
- apply (rule order_trans[OF norm_setsum setsum_mono])
- apply (rule order_trans[OF norm_setsum setsum_mono])
+ apply (simp add: sum_nonneg)
+ apply (simp only: blinfun_of_matrix_apply sum_distrib_right)
+ apply (rule order_trans[OF norm_sum sum_mono])
+ apply (rule order_trans[OF norm_sum sum_mono])
apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm)
done
@@ -419,7 +419,7 @@
have "\<And>i j. i \<in> Basis \<Longrightarrow> j \<in> Basis \<Longrightarrow> Zfun (\<lambda>x. norm (b x i j - a i j)) F"
using assms unfolding tendsto_Zfun_iff Zfun_norm_iff .
hence "Zfun (\<lambda>x. (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>b x i j - a i j\<bar>)) F"
- by (auto intro!: Zfun_setsum)
+ by (auto intro!: Zfun_sum)
thus ?thesis
unfolding tendsto_Zfun_iff blinfun_of_matrix_minus
by (rule Zfun_le) (auto intro!: order_trans[OF norm_blinfun_of_matrix])
@@ -480,8 +480,8 @@
subseq r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) i) (l i) < e) sequentially)"
by (rule compact_lemma_general[where unproj = "\<lambda>e. blinfun_of_matrix (\<lambda>i j. e j \<bullet> i)"])
(auto intro!: euclidean_eqI[where 'a='b] bounded_linear_image assms
- simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta setsum.delta'
- scaleR_setsum_left[symmetric])
+ simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta sum.delta'
+ scaleR_sum_left[symmetric])
lemma blinfun_euclidean_eqI: "(\<And>i. i \<in> Basis \<Longrightarrow> blinfun_apply x i = blinfun_apply y i) \<Longrightarrow> x = y"
apply (auto intro!: blinfun_eqI)
@@ -493,7 +493,7 @@
lemma Blinfun_eq_matrix: "bounded_linear f \<Longrightarrow> Blinfun f = blinfun_of_matrix (\<lambda>i j. f j \<bullet> i)"
by (intro blinfun_euclidean_eqI)
(auto simp: blinfun_of_matrix_apply bounded_linear_Blinfun_apply inner_Basis if_distrib
- cond_application_beta setsum.delta' euclidean_representation
+ cond_application_beta sum.delta' euclidean_representation
cong: if_cong)
text \<open>TODO: generalize (via @{thm compact_cball})?\<close>
@@ -520,10 +520,10 @@
also note norm_blinfun_of_matrix
also have "(\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) <
(\<Sum>i\<in>(Basis::'b set). e / real_of_nat DIM('b))"
- proof (rule setsum_strict_mono)
+ proof (rule sum_strict_mono)
fix i::'b assume i: "i \<in> Basis"
have "(\<Sum>j::'a\<in>Basis. \<bar>(f (r n) - l) j \<bullet> i\<bar>) < (\<Sum>j::'a\<in>Basis. e / ?d)"
- proof (rule setsum_strict_mono)
+ proof (rule sum_strict_mono)
fix j::'a assume j: "j \<in> Basis"
have "\<bar>(f (r n) - l) j \<bullet> i\<bar> \<le> norm ((f (r n) - l) j)"
by (simp add: Basis_le_norm i)
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy Mon Oct 17 14:37:32 2016 +0200
@@ -52,13 +52,13 @@
lemmas Zero_notin_Suc = zero_notin_Suc_image
lemmas atMost_Suc_eq_insert_0 = Iic_Suc_eq_insert_0
-lemma setsum_union_disjoint':
+lemma sum_union_disjoint':
assumes "finite A"
and "finite B"
and "A \<inter> B = {}"
and "A \<union> B = C"
- shows "setsum g C = setsum g A + setsum g B"
- using setsum.union_disjoint[OF assms(1-3)] and assms(4) by auto
+ shows "sum g C = sum g A + sum g B"
+ using sum.union_disjoint[OF assms(1-3)] and assms(4) by auto
lemma pointwise_minimal_pointwise_maximal:
fixes s :: "(nat \<Rightarrow> nat) set"
@@ -139,19 +139,19 @@
shows "odd (card {s\<in>S. compo s})"
proof -
have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + (\<Sum>s | s \<in> S \<and> compo s. nF s) = (\<Sum>s\<in>S. nF s)"
- by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
+ by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong)
also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
(\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
- unfolding setsum.distrib[symmetric]
+ unfolding sum.distrib[symmetric]
by (subst card_Un_disjoint[symmetric])
- (auto simp: nF_def intro!: setsum.cong arg_cong[where f=card])
+ (auto simp: nF_def intro!: sum.cong arg_cong[where f=card])
also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
- using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] setsum_multicount)
+ using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] sum_multicount)
finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
using assms(6,8) by simp
moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) =
(\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 0. nF s) + (\<Sum>s | s \<in> S \<and> \<not> compo s \<and> nF s = 2. nF s)"
- using assms(7) by (subst setsum.union_disjoint[symmetric]) (fastforce intro!: setsum.cong)+
+ using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+
ultimately show ?thesis
by auto
qed
@@ -1417,9 +1417,9 @@
fix y :: 'a assume y: "y \<in> unit_cube"
have "dist 0 y = norm y" by (rule dist_0_norm)
also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation ..
- also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_setsum)
+ also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_sum)
also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)"
- by (rule setsum_mono, simp add: y [unfolded mem_unit_cube])
+ by (rule sum_mono, simp add: y [unfolded mem_unit_cube])
finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" .
qed
@@ -1679,7 +1679,7 @@
unfolding inner_diff_left[symmetric]
by (rule norm_le_l1)
also have "\<dots> < (\<Sum>(i::'a) \<in> Basis. d / real n)"
- apply (rule setsum_strict_mono)
+ apply (rule sum_strict_mono)
using as
apply auto
done
@@ -1731,7 +1731,7 @@
by auto
{
have "(\<Sum>i\<in>Basis. \<bar>real (r (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
- apply (rule setsum_mono)
+ apply (rule sum_mono)
using rs(1)[OF b'_im]
apply (auto simp add:* field_simps simp del: of_nat_Suc)
done
@@ -1743,7 +1743,7 @@
moreover
{
have "(\<Sum>i\<in>Basis. \<bar>real (s (b' i)) - real (q (b' i))\<bar>) \<le> (\<Sum>(i::'a)\<in>Basis. 1)"
- apply (rule setsum_mono)
+ apply (rule sum_mono)
using rs(2)[OF b'_im]
apply (auto simp add:* field_simps simp del: of_nat_Suc)
done
@@ -1757,7 +1757,7 @@
unfolding r'_def s'_def z_def
using \<open>p > 0\<close>
apply (rule_tac[!] le_less_trans[OF norm_le_l1])
- apply (auto simp add: field_simps setsum_divide_distrib[symmetric] inner_diff_left)
+ apply (auto simp add: field_simps sum_divide_distrib[symmetric] inner_diff_left)
done
then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
using rs(3) i
--- a/src/HOL/Analysis/Caratheodory.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Caratheodory.thy Mon Oct 17 14:37:32 2016 +0200
@@ -20,31 +20,31 @@
proof -
have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
using assms by (simp add: fun_eq_iff)
- have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
- by (simp add: setsum.reindex[OF inj_prod_decode] comp_def)
+ have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = sum f (prod_decode ` B)"
+ by (simp add: sum.reindex[OF inj_prod_decode] comp_def)
have "(SUP n. \<Sum>i<n. f (prod_decode i)) = (SUP p : UNIV \<times> UNIV. \<Sum>i<fst p. \<Sum>n<snd p. f (i, n))"
- proof (intro SUP_eq; clarsimp simp: setsum.cartesian_product reindex)
+ proof (intro SUP_eq; clarsimp simp: sum.cartesian_product reindex)
fix n
let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
{ fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
then have "a < ?M fst" "b < ?M snd"
by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
- then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
- by (auto intro!: setsum_mono3)
- then show "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto
+ then have "sum f (prod_decode ` {..<n}) \<le> sum f ({..<?M fst} \<times> {..<?M snd})"
+ by (auto intro!: sum_mono3)
+ then show "\<exists>a b. sum f (prod_decode ` {..<n}) \<le> sum f ({..<a} \<times> {..<b})" by auto
next
fix a b
let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
{ fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
- then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
- by (auto intro!: setsum_mono3)
- then show "\<exists>n. setsum f ({..<a} \<times> {..<b}) \<le> setsum f (prod_decode ` {..<n})"
+ then have "sum f ({..<a} \<times> {..<b}) \<le> sum f ?M"
+ by (auto intro!: sum_mono3)
+ then show "\<exists>n. sum f ({..<a} \<times> {..<b}) \<le> sum f (prod_decode ` {..<n})"
by auto
qed
also have "\<dots> = (SUP p. \<Sum>i<p. \<Sum>n. f (i, n))"
- unfolding suminf_setsum[OF summableI, symmetric]
- by (simp add: suminf_eq_SUP SUP_pair setsum.commute[of _ "{..< fst _}"])
+ unfolding suminf_sum[OF summableI, symmetric]
+ by (simp add: suminf_eq_SUP SUP_pair sum.commute[of _ "{..< fst _}"])
finally show ?thesis unfolding g_def
by (simp add: suminf_eq_SUP)
qed
@@ -529,7 +529,7 @@
with \<open>volume M f\<close> have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
unfolding volume_def by blast
also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
- proof (subst setsum.reindex_nontrivial)
+ proof (subst sum.reindex_nontrivial)
fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j"
with \<open>disjoint_family_on A I\<close> have "A i = {}"
by (auto simp: disjoint_family_on_def)
@@ -547,7 +547,7 @@
shows "volume M \<mu>"
proof (unfold volume_def, safe)
fix C assume "finite C" "C \<subseteq> M" "disjoint C"
- then show "\<mu> (\<Union>C) = setsum \<mu> C"
+ then show "\<mu> (\<Union>C) = sum \<mu> C"
proof (induct C)
case (insert c C)
from insert(1,2,4,5) have "\<mu> (\<Union>insert c C) = \<mu> c + \<mu> (\<Union>C)"
@@ -570,7 +570,7 @@
fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
assume "\<Union>C = \<Union>D"
have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))"
- proof (intro setsum.cong refl)
+ proof (intro sum.cong refl)
fix d assume "d \<in> D"
have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d"
using \<open>d \<in> D\<close> \<open>\<Union>C = \<Union>D\<close> by auto
@@ -592,7 +592,7 @@
assume "\<Union>C = \<Union>D"
with split_sum[OF C D] split_sum[OF D C]
have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
- by (simp, subst setsum.commute, simp add: ac_simps) }
+ by (simp, subst sum.commute, simp add: ac_simps) }
note sum_eq = this
{ fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
@@ -612,7 +612,7 @@
fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
with \<mu>'[of Ca] \<open>volume M \<mu>\<close>[THEN volume_positive]
show "0 \<le> \<mu>' a"
- by (auto intro!: setsum_nonneg)
+ by (auto intro!: sum_nonneg)
next
show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto
next
@@ -627,7 +627,7 @@
also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)"
using C_Int_cases volume_empty[OF \<open>volume M \<mu>\<close>] by (elim disjE) simp_all
also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
- using Ca Cb by (simp add: setsum.union_inter)
+ using Ca Cb by (simp add: sum.union_inter)
also have "\<dots> = \<mu>' a + \<mu>' b"
using Ca Cb by (simp add: \<mu>')
finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b"
@@ -680,7 +680,7 @@
by (simp add: sums_iff)
qed
also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)"
- using F'(2) by (subst (2) F') (simp add: setsum.reindex)
+ using F'(2) by (subst (2) F') (simp add: sum.reindex)
finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" .
next
show "\<mu> {} = 0"
@@ -798,9 +798,9 @@
intro: generated_ringI_Basic)
also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))"
using C' A'
- by (intro suminf_setsum G.Int G.finite_Union) (auto intro: generated_ringI_Basic)
+ by (intro suminf_sum G.Int G.finite_Union) (auto intro: generated_ringI_Basic)
also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
- using eq V C' by (auto intro!: setsum.cong)
+ using eq V C' by (auto intro!: sum.cong)
also have "\<dots> = \<mu>_r (\<Union>C')"
using C' Un_A
by (subst volume_finite_additive[symmetric, OF V(1)])
--- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -12,18 +12,18 @@
by simp
(*move up?*)
-lemma setsum_UNIV_sum:
+lemma sum_UNIV_sum:
fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
apply (subst UNIV_Plus_UNIV [symmetric])
- apply (subst setsum.Plus)
+ apply (subst sum.Plus)
apply simp_all
done
-lemma setsum_mult_product:
- "setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
- unfolding setsum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
-proof (rule setsum.cong, simp, rule setsum.reindex_cong)
+lemma sum_mult_product:
+ "sum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
+ unfolding sum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
+proof (rule sum.cong, simp, rule sum.reindex_cong)
fix i
show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
@@ -108,19 +108,19 @@
subsection \<open>A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space.\<close>
-lemma setsum_cong_aux:
- "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
- by (auto intro: setsum.cong)
+lemma sum_cong_aux:
+ "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> sum f A = sum g A"
+ by (auto intro: sum.cong)
-hide_fact (open) setsum_cong_aux
+hide_fact (open) sum_cong_aux
method_setup vector = \<open>
let
val ss1 =
simpset_of (put_simpset HOL_basic_ss @{context}
- addsimps [@{thm setsum.distrib} RS sym,
- @{thm setsum_subtractf} RS sym, @{thm setsum_distrib_left},
- @{thm setsum_distrib_right}, @{thm setsum_negf} RS sym])
+ addsimps [@{thm sum.distrib} RS sym,
+ @{thm sum_subtractf} RS sym, @{thm sum_distrib_left},
+ @{thm sum_distrib_right}, @{thm sum_negf} RS sym])
val ss2 =
simpset_of (@{context} addsimps
[@{thm plus_vec_def}, @{thm times_vec_def},
@@ -130,8 +130,8 @@
@{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}])
fun vector_arith_tac ctxt ths =
simp_tac (put_simpset ss1 ctxt)
- THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.setsum_cong_aux} i
- ORELSE resolve_tac ctxt @{thms setsum.neutral} i
+ THEN' (fn i => resolve_tac ctxt @{thms Cartesian_Euclidean_Space.sum_cong_aux} i
+ ORELSE resolve_tac ctxt @{thms sum.neutral} i
ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i)
(* THEN' TRY o clarify_tac HOL_cs THEN' (TRY o rtac @{thm iffI}) *)
THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths)
@@ -152,9 +152,9 @@
lemma vec_cmul: "vec(c * x) = c *s vec x " by vector
lemma vec_neg: "vec(- x) = - vec x " by vector
-lemma vec_setsum:
+lemma vec_sum:
assumes "finite S"
- shows "vec(setsum f S) = setsum (vec \<circ> f) S"
+ shows "vec(sum f S) = sum (vec \<circ> f) S"
using assms
proof induct
case empty
@@ -299,8 +299,8 @@
lemma norm_bound_component_lt_cart: "norm x < e ==> \<bar>x$i\<bar> < e"
by (metis component_le_norm_cart le_less_trans)
-lemma norm_le_l1_cart: "norm x <= setsum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
- by (simp add: norm_vec_def setL2_le_setsum)
+lemma norm_le_l1_cart: "norm x <= sum(\<lambda>i. \<bar>x$i\<bar>) UNIV"
+ by (simp add: norm_vec_def setL2_le_sum)
lemma scalar_mult_eq_scaleR: "c *s x = c *\<^sub>R x"
unfolding scaleR_vec_def vector_scalar_mult_def by simp
@@ -309,9 +309,9 @@
unfolding dist_norm scalar_mult_eq_scaleR
unfolding scaleR_right_diff_distrib[symmetric] by simp
-lemma setsum_component [simp]:
+lemma sum_component [simp]:
fixes f:: " 'a \<Rightarrow> ('b::comm_monoid_add) ^'n"
- shows "(setsum f S)$i = setsum (\<lambda>x. (f x)$i) S"
+ shows "(sum f S)$i = sum (\<lambda>x. (f x)$i) S"
proof (cases "finite S")
case True
then show ?thesis by induct simp_all
@@ -320,19 +320,19 @@
then show ?thesis by simp
qed
-lemma setsum_eq: "setsum f S = (\<chi> i. setsum (\<lambda>x. (f x)$i ) S)"
+lemma sum_eq: "sum f S = (\<chi> i. sum (\<lambda>x. (f x)$i ) S)"
by (simp add: vec_eq_iff)
-lemma setsum_cmul:
+lemma sum_cmul:
fixes f:: "'c \<Rightarrow> ('a::semiring_1)^'n"
- shows "setsum (\<lambda>x. c *s f x) S = c *s setsum f S"
- by (simp add: vec_eq_iff setsum_distrib_left)
+ shows "sum (\<lambda>x. c *s f x) S = c *s sum f S"
+ by (simp add: vec_eq_iff sum_distrib_left)
-lemma setsum_norm_allsubsets_bound_cart:
+lemma sum_norm_allsubsets_bound_cart:
fixes f:: "'a \<Rightarrow> real ^'n"
- assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
- shows "setsum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e"
- using setsum_norm_allsubsets_bound[OF assms]
+ assumes fP: "finite P" and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (sum f Q) \<le> e"
+ shows "sum (\<lambda>x. norm (f x)) P \<le> 2 * real CARD('n) * e"
+ using sum_norm_allsubsets_bound[OF assms]
by simp
subsection\<open>Closures and interiors of halfspaces\<close>
@@ -475,15 +475,15 @@
definition matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'p^'n \<Rightarrow> 'a ^ 'p ^'m"
(infixl "**" 70)
- where "m ** m' == (\<chi> i j. setsum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
+ where "m ** m' == (\<chi> i j. sum (\<lambda>k. ((m$i)$k) * ((m'$k)$j)) (UNIV :: 'n set)) ::'a ^ 'p ^'m"
definition matrix_vector_mult :: "('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n \<Rightarrow> 'a ^ 'm"
(infixl "*v" 70)
- where "m *v x \<equiv> (\<chi> i. setsum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
+ where "m *v x \<equiv> (\<chi> i. sum (\<lambda>j. ((m$i)$j) * (x$j)) (UNIV ::'n set)) :: 'a^'m"
definition vector_matrix_mult :: "'a ^ 'm \<Rightarrow> ('a::semiring_1) ^'n^'m \<Rightarrow> 'a ^'n "
(infixl "v*" 70)
- where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
+ where "v v* m == (\<chi> j. sum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
definition transpose where
@@ -495,14 +495,14 @@
lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
- by (vector matrix_matrix_mult_def setsum.distrib[symmetric] field_simps)
+ by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps)
lemma matrix_mul_lid:
fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
shows "mat 1 ** A = A"
apply (simp add: matrix_matrix_mult_def mat_def)
apply vector
- apply (auto simp only: if_distrib cond_application_beta setsum.delta'[OF finite]
+ apply (auto simp only: if_distrib cond_application_beta sum.delta'[OF finite]
mult_1_left mult_zero_left if_True UNIV_I)
done
@@ -512,26 +512,26 @@
shows "A ** mat 1 = A"
apply (simp add: matrix_matrix_mult_def mat_def)
apply vector
- apply (auto simp only: if_distrib cond_application_beta setsum.delta[OF finite]
+ apply (auto simp only: if_distrib cond_application_beta sum.delta[OF finite]
mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
done
lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
- apply (vector matrix_matrix_mult_def setsum_distrib_left setsum_distrib_right mult.assoc)
- apply (subst setsum.commute)
+ apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc)
+ apply (subst sum.commute)
apply simp
done
lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
apply (vector matrix_matrix_mult_def matrix_vector_mult_def
- setsum_distrib_left setsum_distrib_right mult.assoc)
- apply (subst setsum.commute)
+ sum_distrib_left sum_distrib_right mult.assoc)
+ apply (subst sum.commute)
apply simp
done
lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
apply (vector matrix_vector_mult_def mat_def)
- apply (simp add: if_distrib cond_application_beta setsum.delta' cong del: if_weak_cong)
+ apply (simp add: if_distrib cond_application_beta sum.delta' cong del: if_weak_cong)
done
lemma matrix_transpose_mul:
@@ -548,15 +548,15 @@
apply (erule_tac x="axis ia 1" in allE)
apply (erule_tac x="i" in allE)
apply (auto simp add: if_distrib cond_application_beta axis_def
- setsum.delta[OF finite] cong del: if_weak_cong)
+ sum.delta[OF finite] cong del: if_weak_cong)
done
lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \<bullet> x"
by (simp add: matrix_vector_mult_def inner_vec_def)
lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
- apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_distrib_right setsum_distrib_left ac_simps)
- apply (subst setsum.commute)
+ apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps)
+ apply (subst sum.commute)
apply simp
done
@@ -588,27 +588,27 @@
by (simp add: matrix_vector_mult_def inner_vec_def)
lemma matrix_mult_vsum:
- "(A::'a::comm_semiring_1^'n^'m) *v x = setsum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
+ "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\<lambda>i. (x$i) *s column i A) (UNIV:: 'n set)"
by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute)
lemma vector_componentwise:
"(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
- by (simp add: axis_def if_distrib setsum.If_cases vec_eq_iff)
+ by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff)
-lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
- by (auto simp add: axis_def vec_eq_iff if_distrib setsum.If_cases cong del: if_weak_cong)
+lemma basis_expansion: "sum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
+ by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong)
lemma linear_componentwise_expansion:
fixes f:: "real ^'m \<Rightarrow> real ^ _"
assumes lf: "linear f"
- shows "(f x)$j = setsum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
+ shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
proof -
let ?M = "(UNIV :: 'm set)"
let ?N = "(UNIV :: 'n set)"
- have "?rhs = (setsum (\<lambda>i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j"
- unfolding setsum_component by simp
+ have "?rhs = (sum (\<lambda>i.(x$i) *\<^sub>R f (axis i 1) ) ?M)$j"
+ unfolding sum_component by simp
then show ?thesis
- unfolding linear_setsum_mul[OF lf, symmetric]
+ unfolding linear_sum_mul[OF lf, symmetric]
unfolding scalar_mult_eq_scaleR[symmetric]
unfolding basis_expansion
by simp
@@ -630,7 +630,7 @@
lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff
- field_simps setsum_distrib_left setsum.distrib)
+ field_simps sum_distrib_left sum.distrib)
lemma matrix_works:
assumes lf: "linear f"
@@ -652,14 +652,14 @@
by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
lemma matrix_vector_column:
- "(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
+ "(A::'a::comm_semiring_1^'n^_) *v x = sum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute)
lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
apply (rule adjoint_unique)
apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
- setsum_distrib_right setsum_distrib_left)
- apply (subst setsum.commute)
+ sum_distrib_right sum_distrib_left)
+ apply (subst sum.commute)
apply (auto simp add: ac_simps)
done
@@ -758,13 +758,13 @@
lemma matrix_left_invertible_independent_columns:
fixes A :: "real^'n^'m"
shows "(\<exists>(B::real ^'m^'n). B ** A = mat 1) \<longleftrightarrow>
- (\<forall>c. setsum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
+ (\<forall>c. sum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
(is "?lhs \<longleftrightarrow> ?rhs")
proof -
let ?U = "UNIV :: 'n set"
{ assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
{ fix c i
- assume c: "setsum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U"
+ assume c: "sum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U"
let ?x = "\<chi> i. c i"
have th0:"A *v ?x = 0"
using c
@@ -786,7 +786,7 @@
lemma matrix_right_invertible_independent_rows:
fixes A :: "real^'n^'m"
shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow>
- (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
+ (\<forall>c. sum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
unfolding left_invertible_transpose[symmetric]
matrix_left_invertible_independent_columns
by (simp add: column_transpose)
@@ -797,7 +797,7 @@
proof -
let ?U = "UNIV :: 'm set"
have fU: "finite ?U" by simp
- have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y)"
+ have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::real^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)"
unfolding matrix_right_invertible_surjective matrix_mult_vsum surj_def
apply (subst eq_commute)
apply rule
@@ -806,10 +806,10 @@
{ assume h: ?lhs
{ fix x:: "real ^'n"
from h[unfolded lhseq, rule_format, of x] obtain y :: "real ^'m"
- where y: "setsum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
+ where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
have "x \<in> span (columns A)"
unfolding y[symmetric]
- apply (rule span_setsum)
+ apply (rule span_sum)
unfolding scalar_mult_eq_scaleR
apply (rule span_mul)
apply (rule span_superset)
@@ -820,11 +820,11 @@
then have ?rhs unfolding rhseq by blast }
moreover
{ assume h:?rhs
- let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). setsum (\<lambda>i. (x$i) *s column i A) ?U = y"
+ let ?P = "\<lambda>(y::real ^'n). \<exists>(x::real^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y"
{ fix y
have "?P y"
proof (rule span_induct_alt[of ?P "columns A", folded scalar_mult_eq_scaleR])
- show "\<exists>x::real ^ 'm. setsum (\<lambda>i. (x$i) *s column i A) ?U = 0"
+ show "\<exists>x::real ^ 'm. sum (\<lambda>i. (x$i) *s column i A) ?U = 0"
by (rule exI[where x=0], simp)
next
fix c y1 y2
@@ -832,7 +832,7 @@
from y1 obtain i where i: "i \<in> ?U" "y1 = column i A"
unfolding columns_def by blast
from y2 obtain x:: "real ^'m" where
- x: "setsum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
+ x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2" by blast
let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::real^'m"
show "?P (c*s y1 + y2)"
proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left cond_application_beta cong del: if_weak_cong)
@@ -840,18 +840,18 @@
have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))"
using i(1) by (simp add: field_simps)
- have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
- else (x$xa) * ((column xa A$j))) ?U = setsum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
- apply (rule setsum.cong[OF refl])
+ have "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+ else (x$xa) * ((column xa A$j))) ?U = sum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
+ apply (rule sum.cong[OF refl])
using th apply blast
done
- also have "\<dots> = setsum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
- by (simp add: setsum.distrib)
- also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
- unfolding setsum.delta[OF fU]
+ also have "\<dots> = sum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+ by (simp add: sum.distrib)
+ also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
+ unfolding sum.delta[OF fU]
using i(1) by simp
- finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
- else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
+ finally show "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
+ else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
qed
next
show "y \<in> span (columns A)"
@@ -984,9 +984,9 @@
{ fix n
assume n: "\<forall>i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))"
have "dist (f (r n)) l \<le> (\<Sum>i\<in>?d. dist (f (r n) $ i) (l $ i))"
- unfolding dist_vec_def using zero_le_dist by (rule setL2_le_setsum)
+ unfolding dist_vec_def using zero_le_dist by (rule setL2_le_sum)
also have "\<dots> < (\<Sum>i\<in>?d. e / (real_of_nat (card ?d)))"
- by (rule setsum_strict_mono) (simp_all add: n)
+ by (rule sum_strict_mono) (simp_all add: n)
finally have "dist (f (r n)) l < e" by simp
}
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"
@@ -1235,7 +1235,7 @@
where "finite s" "cbox (x - (\<chi> i. d)) (x + (\<chi> i. d)) = convex hull s"
proof -
from assms obtain s where "finite s"
- and "cbox (x - setsum (op *\<^sub>R d) Basis) (x + setsum (op *\<^sub>R d) Basis) = convex hull s"
+ and "cbox (x - sum (op *\<^sub>R d) Basis) (x + sum (op *\<^sub>R d) Basis) = convex hull s"
by (rule cube_convex_hull)
with that[of s] show thesis
by (simp add: const_vector_cart)
@@ -1314,13 +1314,13 @@
lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
using exhaust_3 by auto
-lemma setsum_1: "setsum f (UNIV::1 set) = f 1"
+lemma sum_1: "sum f (UNIV::1 set) = f 1"
unfolding UNIV_1 by simp
-lemma setsum_2: "setsum f (UNIV::2 set) = f 1 + f 2"
+lemma sum_2: "sum f (UNIV::2 set) = f 1 + f 2"
unfolding UNIV_2 by simp
-lemma setsum_3: "setsum f (UNIV::3 set) = f 1 + f 2 + f 3"
+lemma sum_3: "sum f (UNIV::3 set) = f 1 + f 2 + f 3"
unfolding UNIV_3 by (simp add: ac_simps)
instantiation num1 :: cart_one
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1551,9 +1551,9 @@
"(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g"
by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto
-lemma has_contour_integral_setsum:
+lemma has_contour_integral_sum:
"\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk>
- \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p"
+ \<Longrightarrow> ((\<lambda>x. sum (\<lambda>a. f a x) s) has_contour_integral sum i s) p"
by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
@@ -1613,10 +1613,10 @@
lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0"
by (simp add: contour_integral_unique has_contour_integral_0)
-lemma contour_integral_setsum:
+lemma contour_integral_sum:
"\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
- \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s"
- by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral)
+ \<Longrightarrow> contour_integral p (\<lambda>x. sum (\<lambda>a. f a x) s) = sum (\<lambda>a. contour_integral p (f a)) s"
+ by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral)
lemma contour_integrable_eq:
"\<lbrakk>f contour_integrable_on p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g contour_integrable_on p"
@@ -1655,11 +1655,11 @@
using has_contour_integral_div contour_integrable_on_def
by fastforce
-lemma contour_integrable_setsum:
+lemma contour_integrable_sum:
"\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
- \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p"
+ \<Longrightarrow> (\<lambda>x. sum (\<lambda>a. f a x) s) contour_integrable_on p"
unfolding contour_integrable_on_def
- by (metis has_contour_integral_setsum)
+ by (metis has_contour_integral_sum)
subsection\<open>Reversing a path integral\<close>
@@ -5919,9 +5919,9 @@
have sumeq: "(\<Sum>i = 0..n.
of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) =
g z * deriv ((deriv ^^ n) f) z + (\<Sum>i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))"
- apply (simp add: bb algebra_simps setsum.distrib)
- apply (subst (4) setsum_Suc_reindex)
- apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong)
+ apply (simp add: bb algebra_simps sum.distrib)
+ apply (subst (4) sum_Suc_reindex)
+ apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong)
done
show ?case
apply (simp only: funpow.simps o_apply)
@@ -5929,7 +5929,7 @@
apply (rule DERIV_transform_within_open
[of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"])
apply (simp add: algebra_simps)
- apply (rule DERIV_cong [OF DERIV_setsum])
+ apply (rule DERIV_cong [OF DERIV_sum])
apply (rule DERIV_cmult)
apply (auto simp: intro: DERIV_mult * sumeq \<open>open s\<close> Suc.prems Suc.IH [symmetric])
done
@@ -6210,7 +6210,7 @@
by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
= norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)"
- unfolding setsum_distrib_right setsum_divide_distrib power_divide by (simp add: algebra_simps)
+ unfolding sum_distrib_right sum_divide_distrib power_divide by (simp add: algebra_simps)
also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)"
using \<open>0 < B\<close>
apply (auto simp: geometric_sum [OF wzu_not1])
@@ -6242,7 +6242,7 @@
contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
(\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
apply (rule eventuallyI)
- apply (subst contour_integral_setsum, simp)
+ apply (subst contour_integral_sum, simp)
using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps)
apply (simp only: contour_integral_lmul cint algebra_simps)
done
@@ -6251,7 +6251,7 @@
unfolding sums_def
apply (rule Lim_transform_eventually [OF eq])
apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *])
- apply (rule contour_integrable_setsum, simp)
+ apply (rule contour_integrable_sum, simp)
apply (rule contour_integrable_lmul)
apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
using \<open>0 < r\<close>
@@ -6589,7 +6589,7 @@
have "(\<forall>x\<in>s. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> g x)"
using g by (force simp add: lim_sequentially)
moreover have "\<exists>g'. \<forall>x\<in>s. (g has_field_derivative g' x) (at x) \<and> (\<lambda>n. \<Sum>i<n. f' i x) \<longlonglongrightarrow> g' x"
- by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+
+ by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_sum)+
ultimately show ?thesis
by (force simp add: sums_def conj_commute intro: that)
qed
--- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon Oct 17 14:37:32 2016 +0200
@@ -269,7 +269,7 @@
fixes l::complex
shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
unfolding sums_def
-by (metis real_lim_sequentially setsum_in_Reals)
+by (metis real_lim_sequentially sum_in_Reals)
lemma Lim_null_comparison_Re:
assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F"
@@ -335,7 +335,7 @@
"op + c field_differentiable F"
by (simp add: field_differentiable_add)
-lemma field_differentiable_setsum [derivative_intros]:
+lemma field_differentiable_sum [derivative_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i) field_differentiable F) \<Longrightarrow> (\<lambda>z. \<Sum>i\<in>I. f i z) field_differentiable F"
by (induct I rule: infinite_finite_induct)
(auto intro: field_differentiable_add field_differentiable_const)
@@ -502,9 +502,9 @@
"f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_power)
-lemma holomorphic_on_setsum [holomorphic_intros]:
- "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s"
- unfolding holomorphic_on_def by (metis field_differentiable_setsum)
+lemma holomorphic_on_sum [holomorphic_intros]:
+ "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) holomorphic_on s"
+ unfolding holomorphic_on_def by (metis field_differentiable_sum)
lemma DERIV_deriv_iff_field_differentiable:
"DERIV f x :> deriv f x \<longleftrightarrow> f field_differentiable at x"
@@ -788,8 +788,8 @@
"f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s"
by (induct n) (auto simp: analytic_on_const analytic_on_mult)
-lemma analytic_on_setsum:
- "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s"
+lemma analytic_on_sum:
+ "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) analytic_on s"
by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)
lemma deriv_left_inverse:
@@ -951,7 +951,7 @@
then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e"
by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h"
- by (simp add: norm_mult [symmetric] field_simps setsum_distrib_left)
+ by (simp add: norm_mult [symmetric] field_simps sum_distrib_left)
qed
} note ** = this
show ?thesis
@@ -1064,9 +1064,9 @@
subsection \<open>Taylor on Complex Numbers\<close>
-lemma setsum_Suc_reindex:
+lemma sum_Suc_reindex:
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
- shows "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}"
+ shows "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
by (induct n) auto
lemma complex_taylor:
@@ -1146,7 +1146,7 @@
have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)) = (\<Sum>i\<le>n. (f i z / (fact i)) * 0 ^ i)"
by simp
also have "\<dots> = f 0 z / (fact 0)"
- by (subst setsum_zero_power) simp
+ by (subst sum_zero_power) simp
finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)))
\<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) -
(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))"
@@ -1203,7 +1203,7 @@
(\<Sum>i = 0..n.
(f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) /
(fact (Suc i)))"
- by (subst setsum_Suc_reindex) simp
+ by (subst sum_Suc_reindex) simp
also have "... = f (Suc 0) u -
(f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
(fact (Suc n)) +
@@ -1215,7 +1215,7 @@
(f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) /
(fact (Suc n)) +
f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u"
- by (subst setsum_Suc_diff) auto
+ by (subst sum_Suc_diff) auto
also have "... = f (Suc n) u * (z-u) ^ n / (fact n)"
by (simp only: algebra_simps diff_divide_distrib fact_cancel)
finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i
--- a/src/HOL/Analysis/Complex_Transcendental.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Mon Oct 17 14:37:32 2016 +0200
@@ -597,7 +597,7 @@
text\<open>32-bit Approximation to e\<close>
lemma e_approx_32: "\<bar>exp(1) - 5837465777 / 2147483648\<bar> \<le> (inverse(2 ^ 32)::real)"
using Taylor_exp [of 1 14] exp_le
- apply (simp add: setsum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
+ apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
apply (simp only: pos_le_divide_eq [symmetric], linarith)
done
--- a/src/HOL/Analysis/Conformal_Mappings.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Conformal_Mappings.thy Mon Oct 17 14:37:32 2016 +0200
@@ -612,7 +612,7 @@
have "powf sums f w"
unfolding powf_def by (rule holomorphic_power_series [OF holfb w])
moreover have "(\<Sum>i<n. powf i) = f \<xi>"
- apply (subst Groups_Big.comm_monoid_add_class.setsum.setdiff_irrelevant [symmetric])
+ apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric])
apply (simp add:)
apply (simp only: dfz sing)
apply (simp add: powf_def)
@@ -1052,7 +1052,7 @@
proof (cases "\<forall>i\<le>n. 0<i \<longrightarrow> a i = 0")
case True
then have [simp]: "\<And>z. f z = a 0"
- by (simp add: that setsum_atMost_shift)
+ by (simp add: that sum_atMost_shift)
have False using compf [of "{a 0}"] by simp
then show ?thesis ..
next
@@ -1073,7 +1073,7 @@
using k apply auto
done
with k m show ?thesis
- by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.setsum.mono_neutral_right)
+ by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
qed
have "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity"
proof (rule Lim_at_infinityI)
@@ -2758,9 +2758,9 @@
finally show ?thesis .
qed
ultimately show ?case unfolding p_circ_def
- apply (subst (asm) setsum.cong[OF refl,
+ apply (subst (asm) sum.cong[OF refl,
of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"])
- by (auto simp add:setsum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
+ by (auto simp add:sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
qed
lemma Cauchy_theorem_singularities:
@@ -2791,17 +2791,17 @@
show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
by (simp add: avoid pts1_def)
qed
- moreover have "setsum circ pts2=0"
+ moreover have "sum circ pts2=0"
proof -
have "winding_number g p=0" when "p\<in>pts2" for p
using \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto
thus ?thesis unfolding circ_def
- apply (intro setsum.neutral)
+ apply (intro sum.neutral)
by auto
qed
- moreover have "?R=setsum circ pts1 + setsum circ pts2"
+ moreover have "?R=sum circ pts1 + sum circ pts2"
unfolding circ_def
- using setsum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
+ using sum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
by blast
ultimately show ?thesis
apply (fold circ_def)
@@ -2826,7 +2826,7 @@
= (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
using Cauchy_theorem_singularities[OF assms avoid] .
also have "... = (\<Sum>p\<in>pts. c * winding_number g p * residue f p)"
- proof (intro setsum.cong)
+ proof (intro sum.cong)
show "pts = pts" by simp
next
fix x assume "x \<in> pts"
@@ -2846,7 +2846,7 @@
qed
qed
also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
- by (simp add: setsum_distrib_left algebra_simps)
+ by (simp add: sum_distrib_left algebra_simps)
finally show ?thesis unfolding c_def .
qed
@@ -3437,12 +3437,12 @@
qed
also have "... = (\<Sum>p\<in>zeros. w p * cont p) + (\<Sum>p\<in>poles. w p * cont p)"
using finite
- apply (subst setsum.union_disjoint)
+ apply (subst sum.union_disjoint)
by (auto simp add:zeros_def)
also have "... = c * ((\<Sum>p\<in>zeros. w p * h p * zorder f p) - (\<Sum>p\<in>poles. w p * h p * porder f p))"
proof -
have "(\<Sum>p\<in>zeros. w p * cont p) = (\<Sum>p\<in>zeros. c * w p * h p * zorder f p)"
- proof (rule setsum.cong[of zeros zeros,simplified])
+ proof (rule sum.cong[of zeros zeros,simplified])
fix p assume "p \<in> zeros"
show "w p * cont p = c * w p * h p * (zorder f p)"
proof (cases "p\<in>s")
@@ -3459,10 +3459,10 @@
qed
qed
then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros. w p * h p * zorder f p)"
- apply (subst setsum_distrib_left)
+ apply (subst sum_distrib_left)
by (simp add:algebra_simps)
moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles. - c * w p * h p * porder f p)"
- proof (rule setsum.cong[of poles poles,simplified])
+ proof (rule sum.cong[of poles poles,simplified])
fix p assume "p \<in> poles"
show "w p * cont p = - c * w p * h p * (porder f p)"
proof (cases "p\<in>s")
@@ -3479,7 +3479,7 @@
qed
qed
then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p * h p * porder f p)"
- apply (subst setsum_distrib_left)
+ apply (subst sum_distrib_left)
by (simp add:algebra_simps)
ultimately show ?thesis by (simp add: right_diff_distrib)
qed
--- a/src/HOL/Analysis/Continuous_Extension.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Continuous_Extension.thy Mon Oct 17 14:37:32 2016 +0200
@@ -20,19 +20,19 @@
obtains F :: "['a set, 'a] \<Rightarrow> real"
where "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x \<in> S. 0 \<le> F U x)"
and "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
- and "\<And>x. x \<in> S \<Longrightarrow> supp_setsum (\<lambda>W. F W x) \<C> = 1"
+ and "\<And>x. x \<in> S \<Longrightarrow> supp_sum (\<lambda>W. F W x) \<C> = 1"
and "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}"
proof (cases "\<exists>W. W \<in> \<C> \<and> S \<subseteq> W")
case True
then obtain W where "W \<in> \<C>" "S \<subseteq> W" by metis
then show ?thesis
apply (rule_tac F = "\<lambda>V x. if V = W then 1 else 0" in that)
- apply (auto simp: continuous_on_const supp_setsum_def support_on_def)
+ apply (auto simp: continuous_on_const supp_sum_def support_on_def)
done
next
case False
- have nonneg: "0 \<le> supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" for x
- by (simp add: supp_setsum_def setsum_nonneg)
+ have nonneg: "0 \<le> supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>" for x
+ by (simp add: supp_sum_def sum_nonneg)
have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> V" for V x
proof -
have "closedin (subtopology euclidean S) (S - V)"
@@ -41,7 +41,7 @@
show ?thesis
by (simp add: order_class.order.order_iff_strict)
qed
- have ss_pos: "0 < supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x
+ have ss_pos: "0 < supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x
proof -
obtain U where "U \<in> \<C>" "x \<in> U" using \<open>x \<in> S\<close> \<open>S \<subseteq> \<Union>\<C>\<close>
by blast
@@ -52,20 +52,20 @@
have "x \<notin> closure (S - U)"
by (metis \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> less_irrefl sd_pos setdist_eq_0_sing_1 that)
then show ?thesis
- apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def)
- apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U])
+ apply (simp add: setdist_eq_0_sing_1 supp_sum_def support_on_def)
+ apply (rule ordered_comm_monoid_add_class.sum_pos2 [OF *, of U])
using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> False
apply (auto simp: setdist_pos_le sd_pos that)
done
qed
define F where
- "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>
+ "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>
else 0"
show ?thesis
proof (rule_tac F = F in that)
have "continuous_on S (F U)" if "U \<in> \<C>" for U
proof -
- have *: "continuous_on S (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
+ have *: "continuous_on S (\<lambda>x. supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>)"
proof (clarsimp simp add: continuous_on_eq_continuous_within)
fix x assume "x \<in> S"
then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}"
@@ -73,15 +73,15 @@
then have OSX: "openin (subtopology euclidean S) (S \<inter> X)" by blast
have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow>
(\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V))
- = supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>"
- apply (simp add: supp_setsum_def)
- apply (rule setsum.mono_neutral_right [OF finX])
+ = supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>"
+ apply (simp add: supp_sum_def)
+ apply (rule sum.mono_neutral_right [OF finX])
apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff)
apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE)
done
- show "continuous (at x within S) (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
+ show "continuous (at x within S) (\<lambda>x. supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>)"
apply (rule continuous_transform_within_openin
- [where f = "\<lambda>x. (setsum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})"
+ [where f = "\<lambda>x. (sum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})"
and S ="S \<inter> X"])
apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+
apply (simp add: sumeq)
@@ -101,9 +101,9 @@
show "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
by (simp add: setdist_eq_0_sing_1 closure_def F_def)
next
- show "supp_setsum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x
+ show "supp_sum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x
using that ss_pos [OF that]
- by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric])
+ by (simp add: F_def divide_simps supp_sum_divide_distrib [symmetric])
next
show "\<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" if "x \<in> S" for x
using fin [OF that] that
@@ -378,11 +378,11 @@
where Hcont: "\<And>Z. Z \<in> \<C> \<Longrightarrow> continuous_on (U-S) (H Z)"
and Hge0: "\<And>Z x. \<lbrakk>Z \<in> \<C>; x \<in> U-S\<rbrakk> \<Longrightarrow> 0 \<le> H Z x"
and Heq0: "\<And>x Z. \<lbrakk>Z \<in> \<C>; x \<in> U-S; x \<notin> Z\<rbrakk> \<Longrightarrow> H Z x = 0"
- and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_setsum (\<lambda>W. H W x) \<C> = 1"
+ and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_sum (\<lambda>W. H W x) \<C> = 1"
and Hfin: "\<And>x. x \<in> U-S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. H U x \<noteq> 0}"
apply (rule subordinate_partition_of_unity [OF USsub _ fin])
using nbrhd by auto
- define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_setsum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>"
+ define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_sum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>"
show ?thesis
proof (rule that)
show "continuous_on U g"
@@ -420,8 +420,8 @@
then show ?thesis
using VA [OF \<open>T \<in> \<C>\<close>] by (auto simp: d)
qed
- have "supp_setsum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e"
- apply (rule convex_supp_setsum [OF convex_ball])
+ have "supp_sum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e"
+ apply (rule convex_supp_sum [OF convex_ball])
apply (simp_all add: False H1 Hge0 \<open>y \<in> U\<close>)
by (metis dist_commute *)
then show ?thesis
@@ -447,7 +447,7 @@
done
show ?thesis
proof (rule continuous_transform_within_openin [OF _ oUS])
- show "continuous (at a within U) (\<lambda>x. supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)"
+ show "continuous (at a within U) (\<lambda>x. supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)"
proof (rule continuous_transform_within_openin)
show "continuous (at a within U)
(\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))"
@@ -460,21 +460,21 @@
next
show "\<And>x. x \<in> (U - S) \<inter> N \<Longrightarrow>
(\<Sum>T \<in> {U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))
- = supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>"
- by (auto simp: supp_setsum_def support_on_def
- intro: setsum.mono_neutral_right [OF finN])
+ = supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>"
+ by (auto simp: supp_sum_def support_on_def
+ intro: sum.mono_neutral_right [OF finN])
qed
next
show "a \<in> U - S" using False \<open>a \<in> U\<close> by blast
next
- show "\<And>x. x \<in> U - S \<Longrightarrow> supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x"
+ show "\<And>x. x \<in> U - S \<Longrightarrow> supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x"
by (simp add: g_def)
qed
qed
qed
show "g ` U \<subseteq> C"
using \<open>f ` S \<subseteq> C\<close> VA
- by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \<open>convex C\<close>] H1)
+ by (fastforce simp: g_def Hge0 intro!: convex_supp_sum [OF \<open>convex C\<close>] H1)
show "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
by (simp add: g_def)
qed
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -69,9 +69,9 @@
by (auto intro: finite_subset[OF assms])
have ***: "\<And>i. i \<in> Basis \<Longrightarrow> (\<Sum>i\<in>d. f i *\<^sub>R i) \<bullet> i = (\<Sum>x\<in>d. if x = i then f x else 0)"
using d
- by (auto intro!: setsum.cong simp: inner_Basis inner_setsum_left)
- show ?thesis
- unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum.delta[OF **] ***)
+ by (auto intro!: sum.cong simp: inner_Basis inner_sum_left)
+ show ?thesis
+ unfolding euclidean_eq_iff[where 'a='a] by (auto simp: sum.delta[OF **] ***)
qed
lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
@@ -106,7 +106,7 @@
shows "finite B \<and> card B = dim (span B)"
using assms basis_card_eq_dim[of B "span B"] span_inc by auto
-lemma setsum_not_0: "setsum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
+lemma sum_not_0: "sum f A \<noteq> 0 \<Longrightarrow> \<exists>a \<in> A. f a \<noteq> 0"
by (rule ccontr) auto
lemma subset_translation_eq [simp]:
@@ -306,7 +306,7 @@
subsection \<open>Explicit expressions for convexity in terms of arbitrary sums\<close>
-lemma convex_setsum:
+lemma convex_sum:
fixes C :: "'a::real_vector set"
assumes "finite s"
and "convex C"
@@ -320,55 +320,55 @@
then show ?case by simp
next
case (insert i s) note IH = this(3)
- have "a i + setsum a s = 1"
+ have "a i + sum a s = 1"
and "0 \<le> a i"
and "\<forall>j\<in>s. 0 \<le> a j"
and "y i \<in> C"
and "\<forall>j\<in>s. y j \<in> C"
using insert.hyps(1,2) insert.prems by simp_all
- then have "0 \<le> setsum a s"
- by (simp add: setsum_nonneg)
+ then have "0 \<le> sum a s"
+ by (simp add: sum_nonneg)
have "a i *\<^sub>R y i + (\<Sum>j\<in>s. a j *\<^sub>R y j) \<in> C"
- proof (cases "setsum a s = 0")
+ proof (cases "sum a s = 0")
case True
- with \<open>a i + setsum a s = 1\<close> have "a i = 1"
+ with \<open>a i + sum a s = 1\<close> have "a i = 1"
by simp
- from setsum_nonneg_0 [OF \<open>finite s\<close> _ True] \<open>\<forall>j\<in>s. 0 \<le> a j\<close> have "\<forall>j\<in>s. a j = 0"
+ from sum_nonneg_0 [OF \<open>finite s\<close> _ True] \<open>\<forall>j\<in>s. 0 \<le> a j\<close> have "\<forall>j\<in>s. a j = 0"
by simp
show ?thesis using \<open>a i = 1\<close> and \<open>\<forall>j\<in>s. a j = 0\<close> and \<open>y i \<in> C\<close>
by simp
next
case False
- with \<open>0 \<le> setsum a s\<close> have "0 < setsum a s"
+ with \<open>0 \<le> sum a s\<close> have "0 < sum a s"
by simp
- then have "(\<Sum>j\<in>s. (a j / setsum a s) *\<^sub>R y j) \<in> C"
+ then have "(\<Sum>j\<in>s. (a j / sum a s) *\<^sub>R y j) \<in> C"
using \<open>\<forall>j\<in>s. 0 \<le> a j\<close> and \<open>\<forall>j\<in>s. y j \<in> C\<close>
- by (simp add: IH setsum_divide_distrib [symmetric])
+ by (simp add: IH sum_divide_distrib [symmetric])
from \<open>convex C\<close> and \<open>y i \<in> C\<close> and this and \<open>0 \<le> a i\<close>
- and \<open>0 \<le> setsum a s\<close> and \<open>a i + setsum a s = 1\<close>
- have "a i *\<^sub>R y i + setsum a s *\<^sub>R (\<Sum>j\<in>s. (a j / setsum a s) *\<^sub>R y j) \<in> C"
+ and \<open>0 \<le> sum a s\<close> and \<open>a i + sum a s = 1\<close>
+ have "a i *\<^sub>R y i + sum a s *\<^sub>R (\<Sum>j\<in>s. (a j / sum a s) *\<^sub>R y j) \<in> C"
by (rule convexD)
then show ?thesis
- by (simp add: scaleR_setsum_right False)
+ by (simp add: scaleR_sum_right False)
qed
then show ?case using \<open>finite s\<close> and \<open>i \<notin> s\<close>
by simp
qed
lemma convex:
- "convex s \<longleftrightarrow> (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>s) \<and> (setsum u {1..k} = 1)
- \<longrightarrow> setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
+ "convex s \<longleftrightarrow> (\<forall>(k::nat) u x. (\<forall>i. 1\<le>i \<and> i\<le>k \<longrightarrow> 0 \<le> u i \<and> x i \<in>s) \<and> (sum u {1..k} = 1)
+ \<longrightarrow> sum (\<lambda>i. u i *\<^sub>R x i) {1..k} \<in> s)"
proof safe
fix k :: nat
fix u :: "nat \<Rightarrow> real"
fix x
assume "convex s"
"\<forall>i. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s"
- "setsum u {1..k} = 1"
- with convex_setsum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
- by auto
-next
- assume *: "\<forall>k u x. (\<forall> i :: nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1
+ "sum u {1..k} = 1"
+ with convex_sum[of "{1 .. k}" s] show "(\<Sum>j\<in>{1 .. k}. u j *\<^sub>R x j) \<in> s"
+ by auto
+next
+ assume *: "\<forall>k u x. (\<forall> i :: nat. 1 \<le> i \<and> i \<le> k \<longrightarrow> 0 \<le> u i \<and> x i \<in> s) \<and> sum u {1..k} = 1
\<longrightarrow> (\<Sum>i = 1..k. u i *\<^sub>R (x i :: 'a)) \<in> s"
{
fix \<mu> :: real
@@ -381,14 +381,14 @@
by auto
then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
by simp
- then have "setsum ?u {1 .. 2} = 1"
- using setsum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
+ then have "sum ?u {1 .. 2} = 1"
+ using sum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
by auto
with *[rule_format, of "2" ?u ?x] have s: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> s"
using mu xy by auto
have grarr: "(\<Sum>j \<in> {Suc (Suc 0)..2}. ?u j *\<^sub>R ?x j) = (1 - \<mu>) *\<^sub>R y"
- using setsum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto
- from setsum_head_Suc[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this]
+ using sum_head_Suc[of "Suc (Suc 0)" 2 "\<lambda> j. (1 - \<mu>) *\<^sub>R y"] by auto
+ from sum_head_Suc[of "Suc 0" 2 "\<lambda> j. ?u j *\<^sub>R ?x j", simplified this]
have "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) = \<mu> *\<^sub>R x + (1 - \<mu>) *\<^sub>R y"
by auto
then have "(1 - \<mu>) *\<^sub>R y + \<mu> *\<^sub>R x \<in> s"
@@ -402,18 +402,18 @@
lemma convex_explicit:
fixes s :: "'a::real_vector set"
shows "convex s \<longleftrightarrow>
- (\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> setsum u t = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) t \<in> s)"
+ (\<forall>t u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and> sum u t = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) t \<in> s)"
proof safe
fix t
fix u :: "'a \<Rightarrow> real"
assume "convex s"
and "finite t"
- and "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1"
+ and "t \<subseteq> s" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1"
then show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
- using convex_setsum[of t s u "\<lambda> x. x"] by auto
+ using convex_sum[of t s u "\<lambda> x. x"] by auto
next
assume *: "\<forall>t. \<forall> u. finite t \<and> t \<subseteq> s \<and> (\<forall>x\<in>t. 0 \<le> u x) \<and>
- setsum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
+ sum u t = 1 \<longrightarrow> (\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
show "convex s"
unfolding convex_alt
proof safe
@@ -437,7 +437,7 @@
lemma convex_finite:
assumes "finite s"
- shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> setsum (\<lambda>x. u x *\<^sub>R x) s \<in> s)"
+ shows "convex s \<longleftrightarrow> (\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<longrightarrow> sum (\<lambda>x. u x *\<^sub>R x) s \<in> s)"
unfolding convex_explicit
apply safe
subgoal for u by (erule allE [where x=s], erule allE [where x=u]) auto
@@ -445,12 +445,12 @@
proof -
have if_distrib_arg: "\<And>P f g x. (if P then f else g) x = (if P then f x else g x)"
by simp
- assume sum: "\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s"
- assume *: "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1"
+ assume sum: "\<forall>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> s"
+ assume *: "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1"
assume "t \<subseteq> s"
then have "s \<inter> t = t" by auto
with sum[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]] * show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
- by (auto simp: assms setsum.If_cases if_distrib if_distrib_arg)
+ by (auto simp: assms sum.If_cases if_distrib if_distrib_arg)
qed
done
@@ -704,7 +704,7 @@
by fastforce
qed
-lemma convex_on_setsum:
+lemma convex_on_sum:
fixes a :: "'a \<Rightarrow> real"
and y :: "'a \<Rightarrow> 'b::real_vector"
and f :: "'b \<Rightarrow> real"
@@ -736,7 +736,7 @@
then have "(\<Sum> j \<in> s. a j) = 0"
using insert by auto
then have "\<And>j. j \<in> s \<Longrightarrow> a j = 0"
- using insert by (fastforce simp: setsum_nonneg_eq_0_iff)
+ using insert by (fastforce simp: sum_nonneg_eq_0_iff)
then show ?thesis
using insert by auto
next
@@ -746,7 +746,7 @@
have fis: "finite (insert i s)"
using insert by auto
then have ai1: "a i \<le> 1"
- using setsum_nonneg_leq_bound[of "insert i s" a] insert by simp
+ using sum_nonneg_leq_bound[of "insert i s" a] insert by simp
then have "a i < 1"
using False by auto
then have i0: "1 - a i > 0"
@@ -757,23 +757,23 @@
have "(\<Sum> j \<in> insert i s. a j) = 1"
using insert by auto
then have "(\<Sum> j \<in> s. a j) = 1 - a i"
- using setsum.insert insert by fastforce
+ using sum.insert insert by fastforce
then have "(\<Sum> j \<in> s. a j) / (1 - a i) = 1"
using i0 by auto
then have a1: "(\<Sum> j \<in> s. ?a j) = 1"
- unfolding setsum_divide_distrib by simp
+ unfolding sum_divide_distrib by simp
have "convex C" using insert by auto
then have asum: "(\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<in> C"
- using insert convex_setsum [OF \<open>finite s\<close> \<open>convex C\<close> a1 a_nonneg] by auto
+ using insert convex_sum [OF \<open>finite s\<close> \<open>convex C\<close> a1 a_nonneg] by auto
have asum_le: "f (\<Sum> j \<in> s. ?a j *\<^sub>R y j) \<le> (\<Sum> j \<in> s. ?a j * f (y j))"
using a_nonneg a1 insert by blast
have "f (\<Sum> j \<in> insert i s. a j *\<^sub>R y j) = f ((\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
- using setsum.insert[of s i "\<lambda> j. a j *\<^sub>R y j", OF \<open>finite s\<close> \<open>i \<notin> s\<close>] insert
+ using sum.insert[of s i "\<lambda> j. a j *\<^sub>R y j", OF \<open>finite s\<close> \<open>i \<notin> s\<close>] insert
by (auto simp only: add.commute)
also have "\<dots> = f (((1 - a i) * inverse (1 - a i)) *\<^sub>R (\<Sum> j \<in> s. a j *\<^sub>R y j) + a i *\<^sub>R y i)"
using i0 by auto
also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. (a j * inverse (1 - a i)) *\<^sub>R y j) + a i *\<^sub>R y i)"
- using scaleR_right.setsum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric]
+ using scaleR_right.sum[of "inverse (1 - a i)" "\<lambda> j. a j *\<^sub>R y j" s, symmetric]
by (auto simp: algebra_simps)
also have "\<dots> = f ((1 - a i) *\<^sub>R (\<Sum> j \<in> s. ?a j *\<^sub>R y j) + a i *\<^sub>R y i)"
by (auto simp: divide_inverse)
@@ -785,7 +785,7 @@
OF asum_le less_imp_le[OF i0]], of "a i * f (y i)"]
by simp
also have "\<dots> = (\<Sum> j \<in> s. (1 - a i) * ?a j * f (y j)) + a i * f (y i)"
- unfolding setsum_distrib_left[of "1 - a i" "\<lambda> j. ?a j * f (y j)"]
+ unfolding sum_distrib_left[of "1 - a i" "\<lambda> j. ?a j * f (y j)"]
using i0 by auto
also have "\<dots> = (\<Sum> j \<in> s. a j * f (y j)) + a i * f (y i)"
using i0 by auto
@@ -1148,19 +1148,19 @@
finally show ?thesis .
qed (insert assms(2), simp_all)
-lemma convex_supp_setsum:
- assumes "convex S" and 1: "supp_setsum u I = 1"
+lemma convex_supp_sum:
+ assumes "convex S" and 1: "supp_sum u I = 1"
and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)"
- shows "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
+ shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
proof -
have fin: "finite {i \<in> I. u i \<noteq> 0}"
- using 1 setsum.infinite by (force simp: supp_setsum_def support_on_def)
- then have eq: "supp_setsum (\<lambda>i. u i *\<^sub>R f i) I = setsum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
- by (force intro: setsum.mono_neutral_left simp: supp_setsum_def support_on_def)
+ using 1 sum.infinite by (force simp: supp_sum_def support_on_def)
+ then have eq: "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
+ by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def)
show ?thesis
apply (simp add: eq)
- apply (rule convex_setsum [OF fin \<open>convex S\<close>])
- using 1 assms apply (auto simp: supp_setsum_def support_on_def)
+ apply (rule convex_sum [OF fin \<open>convex S\<close>])
+ using 1 assms apply (auto simp: supp_sum_def support_on_def)
done
qed
@@ -1335,18 +1335,18 @@
shows "sphere a r = {} \<longleftrightarrow> r < 0"
by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist)
-lemma setsum_delta_notmem:
+lemma sum_delta_notmem:
assumes "x \<notin> s"
- shows "setsum (\<lambda>y. if (y = x) then P x else Q y) s = setsum Q s"
- and "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
- and "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
- and "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
- apply (rule_tac [!] setsum.cong)
+ shows "sum (\<lambda>y. if (y = x) then P x else Q y) s = sum Q s"
+ and "sum (\<lambda>y. if (x = y) then P x else Q y) s = sum Q s"
+ and "sum (\<lambda>y. if (y = x) then P y else Q y) s = sum Q s"
+ and "sum (\<lambda>y. if (x = y) then P y else Q y) s = sum Q s"
+ apply (rule_tac [!] sum.cong)
using assms
apply auto
done
-lemma setsum_delta'':
+lemma sum_delta'':
fixes s::"'a::real_vector set"
assumes "finite s"
shows "(\<Sum>x\<in>s. (if y = x then f x else 0) *\<^sub>R x) = (if y\<in>s then (f y) *\<^sub>R y else 0)"
@@ -1354,7 +1354,7 @@
have *: "\<And>x y. (if y = x then f x else (0::real)) *\<^sub>R x = (if x=y then (f x) *\<^sub>R x else 0)"
by auto
show ?thesis
- unfolding * using setsum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
+ unfolding * using sum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
qed
lemma if_smult: "(if P then x else (y::real)) *\<^sub>R v = (if P then x *\<^sub>R v else y *\<^sub>R v)"
@@ -1429,7 +1429,7 @@
lemma affine:
fixes V::"'a::real_vector set"
shows "affine V \<longleftrightarrow>
- (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (setsum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
+ (\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (sum (\<lambda>x. (u x) *\<^sub>R x)) s \<in> V)"
unfolding affine_def
apply rule
apply(rule, rule, rule)
@@ -1439,7 +1439,7 @@
proof -
fix x y u v
assume as: "x \<in> V" "y \<in> V" "u + v = (1::real)"
- "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> setsum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
+ "\<forall>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> V \<and> sum u s = 1 \<longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
then show "u *\<^sub>R x + v *\<^sub>R y \<in> V"
apply (cases "x = y")
using as(4)[THEN spec[where x="{x,y}"], THEN spec[where x="\<lambda>w. if w = x then u else v"]]
@@ -1449,7 +1449,7 @@
next
fix s u
assume as: "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
- "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = (1::real)"
+ "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = (1::real)"
define n where "n = card s"
have "card s = 0 \<or> card s = 1 \<or> card s = 2 \<or> card s > 2" by auto
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
@@ -1461,7 +1461,7 @@
unfolding card_Suc_eq by auto
then show ?thesis
using as(1)[THEN bspec[where x=a], THEN bspec[where x=b]] using as(4,5)
- by (auto simp add: setsum_clauses(2))
+ by (auto simp add: sum_clauses(2))
next
assume "card s > 2"
then show ?thesis using as and n_def
@@ -1473,15 +1473,15 @@
fix s :: "'a set" and u :: "'a \<Rightarrow> real"
assume IA:
"\<And>u s. \<lbrakk>2 < card s; \<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V; finite s;
- s \<noteq> {}; s \<subseteq> V; setsum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
+ s \<noteq> {}; s \<subseteq> V; sum u s = 1; n = card s \<rbrakk> \<Longrightarrow> (\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
and as:
"Suc n = card s" "2 < card s" "\<forall>x\<in>V. \<forall>y\<in>V. \<forall>u v. u + v = 1 \<longrightarrow> u *\<^sub>R x + v *\<^sub>R y \<in> V"
- "finite s" "s \<noteq> {}" "s \<subseteq> V" "setsum u s = 1"
+ "finite s" "s \<noteq> {}" "s \<subseteq> V" "sum u s = 1"
have "\<exists>x\<in>s. u x \<noteq> 1"
proof (rule ccontr)
assume "\<not> ?thesis"
- then have "setsum u s = real_of_nat (card s)"
- unfolding card_eq_setsum by auto
+ then have "sum u s = real_of_nat (card s)"
+ unfolding card_eq_sum by auto
then show False
using as(7) and \<open>card s > 2\<close>
by (metis One_nat_def less_Suc0 Zero_not_Suc of_nat_1 of_nat_eq_iff numeral_2_eq_2)
@@ -1495,9 +1495,9 @@
done
have *: "s = insert x (s - {x})" "finite (s - {x})"
using \<open>x\<in>s\<close> and as(4) by auto
- have **: "setsum u (s - {x}) = 1 - u x"
- using setsum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
- have ***: "inverse (1 - u x) * setsum u (s - {x}) = 1"
+ have **: "sum u (s - {x}) = 1 - u x"
+ using sum_clauses(2)[OF *(2), of u x, unfolded *(1)[symmetric] as(7)] by auto
+ have ***: "inverse (1 - u x) * sum u (s - {x}) = 1"
unfolding ** using \<open>u x \<noteq> 1\<close> by auto
have "(\<Sum>xa\<in>s - {x}. (inverse (1 - u x) * u xa) *\<^sub>R xa) \<in> V"
proof (cases "card (s - {x}) > 2")
@@ -1511,7 +1511,7 @@
qed auto
then show ?thesis
apply (rule_tac IA[of "s - {x}" "\<lambda>y. (inverse (1 - u x) * u y)"])
- unfolding setsum_distrib_left[symmetric]
+ unfolding sum_distrib_left[symmetric]
using as and *** and True
apply auto
done
@@ -1524,21 +1524,21 @@
then show ?thesis
using as(3)[THEN bspec[where x=a], THEN bspec[where x=b]]
using *** *(2) and \<open>s \<subseteq> V\<close>
- unfolding setsum_distrib_left
- by (auto simp add: setsum_clauses(2))
+ unfolding sum_distrib_left
+ by (auto simp add: sum_clauses(2))
qed
then have "u x + (1 - u x) = 1 \<Longrightarrow>
u x *\<^sub>R x + (1 - u x) *\<^sub>R ((\<Sum>xa\<in>s - {x}. u xa *\<^sub>R xa) /\<^sub>R (1 - u x)) \<in> V"
apply -
apply (rule as(3)[rule_format])
- unfolding Real_Vector_Spaces.scaleR_right.setsum
+ unfolding Real_Vector_Spaces.scaleR_right.sum
using x(1) as(6)
apply auto
done
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> V"
- unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
+ unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric]
apply (subst *)
- unfolding setsum_clauses(2)[OF *(2)]
+ unfolding sum_clauses(2)[OF *(2)]
using \<open>u x \<noteq> 1\<close>
apply auto
done
@@ -1554,7 +1554,7 @@
lemma affine_hull_explicit:
"affine hull p =
- {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> setsum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
+ {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> sum (\<lambda>v. (u v) *\<^sub>R v) s = y}"
apply (rule hull_unique)
apply (subst subset_eq)
prefer 3
@@ -1567,7 +1567,7 @@
proof -
fix x
assume "x\<in>p"
- then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ then show "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
apply (rule_tac x="{x}" in exI)
apply (rule_tac x="\<lambda>x. 1" in exI)
apply auto
@@ -1575,12 +1575,12 @@
next
fix t x s u
assume as: "p \<subseteq> t" "affine t" "finite s" "s \<noteq> {}"
- "s \<subseteq> p" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ "s \<subseteq> p" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
then show "x \<in> t"
using as(2)[unfolded affine, THEN spec[where x=s], THEN spec[where x=u]]
by auto
next
- show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
+ show "affine {y. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y}"
unfolding affine_def
apply (rule, rule, rule, rule, rule)
unfolding mem_Collect_eq
@@ -1588,26 +1588,26 @@
fix u v :: real
assume uv: "u + v = 1"
fix x
- assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
then obtain sx ux where
- x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "setsum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
+ x: "finite sx" "sx \<noteq> {}" "sx \<subseteq> p" "sum ux sx = 1" "(\<Sum>v\<in>sx. ux v *\<^sub>R v) = x"
by auto
fix y
- assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+ assume "\<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
then obtain sy uy where
- y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "setsum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
+ y: "finite sy" "sy \<noteq> {}" "sy \<subseteq> p" "sum uy sy = 1" "(\<Sum>v\<in>sy. uy v *\<^sub>R v) = y" by auto
have xy: "finite (sx \<union> sy)"
using x(1) y(1) by auto
have **: "(sx \<union> sy) \<inter> sx = sx" "(sx \<union> sy) \<inter> sy = sy"
by auto
show "\<exists>s ua. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p \<and>
- setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
+ sum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
apply (rule_tac x="sx \<union> sy" in exI)
apply (rule_tac x="\<lambda>a. (if a\<in>sx then u * ux a else 0) + (if a\<in>sy then v * uy a else 0)" in exI)
- unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left
- ** setsum.inter_restrict[OF xy, symmetric]
- unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric]
- and setsum_distrib_left[symmetric]
+ unfolding scaleR_left_distrib sum.distrib if_smult scaleR_zero_left
+ ** sum.inter_restrict[OF xy, symmetric]
+ unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric]
+ and sum_distrib_left[symmetric]
unfolding x y
using x(1-3) y(1-3) uv
apply simp
@@ -1617,7 +1617,7 @@
lemma affine_hull_finite:
assumes "finite s"
- shows "affine hull s = {y. \<exists>u. setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+ shows "affine hull s = {y. \<exists>u. sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
unfolding affine_hull_explicit and set_eq_iff and mem_Collect_eq
apply (rule, rule)
apply (erule exE)+
@@ -1627,9 +1627,9 @@
apply (erule conjE)
proof -
fix x u
- assume "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ assume "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
then show "\<exists>sa u. finite sa \<and>
- \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
+ \<not> (\<forall>x. (x \<in> sa) = (x \<in> {})) \<and> sa \<subseteq> s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = x"
apply (rule_tac x=s in exI, rule_tac x=u in exI)
using assms
apply auto
@@ -1639,10 +1639,10 @@
assume "t \<subseteq> s"
then have *: "s \<inter> t = t"
by auto
- assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
- then show "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ assume "finite t" "\<not> (\<forall>x. (x \<in> t) = (x \<in> {}))" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
+ then show "\<exists>u. sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
- unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms, symmetric] and *
+ unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms, symmetric] and *
apply auto
done
qed
@@ -1653,22 +1653,22 @@
lemma affine_hull_empty[simp]: "affine hull {} = {}"
by (rule hull_unique) auto
-(*could delete: it simply rewrites setsum expressions, but it's used twice*)
+(*could delete: it simply rewrites sum expressions, but it's used twice*)
lemma affine_hull_finite_step:
fixes y :: "'a::real_vector"
shows
- "(\<exists>u. setsum u {} = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
+ "(\<exists>u. sum u {} = w \<and> sum (\<lambda>x. u x *\<^sub>R x) {} = y) \<longleftrightarrow> w = 0 \<and> y = 0" (is ?th1)
and
"finite s \<Longrightarrow>
- (\<exists>u. setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
- (\<exists>v u. setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
+ (\<exists>u. sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y) \<longleftrightarrow>
+ (\<exists>v u. sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)" (is "_ \<Longrightarrow> ?lhs = ?rhs")
proof -
show ?th1 by simp
assume fin: "finite s"
show "?lhs = ?rhs"
proof
assume ?lhs
- then obtain u where u: "setsum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
+ then obtain u where u: "sum u (insert a s) = w \<and> (\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
by auto
show ?rhs
proof (cases "a \<in> s")
@@ -1689,7 +1689,7 @@
qed
next
assume ?rhs
- then obtain v u where vu: "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+ then obtain v u where vu: "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
by auto
have *: "\<And>x M. (if x = a then v else M) *\<^sub>R x = (if x = a then v *\<^sub>R x else M *\<^sub>R x)"
by auto
@@ -1698,11 +1698,11 @@
case True
then show ?thesis
apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
- unfolding setsum_clauses(2)[OF fin]
+ unfolding sum_clauses(2)[OF fin]
apply simp
- unfolding scaleR_left_distrib and setsum.distrib
+ unfolding scaleR_left_distrib and sum.distrib
unfolding vu and * and scaleR_zero_left
- apply (auto simp add: setsum.delta[OF fin])
+ apply (auto simp add: sum.delta[OF fin])
done
next
case False
@@ -1711,9 +1711,9 @@
"\<And>x. x \<in> s \<Longrightarrow> u x *\<^sub>R x = (if x = a then v *\<^sub>R x else u x *\<^sub>R x)" by auto
from False show ?thesis
apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
- unfolding setsum_clauses(2)[OF fin] and * using vu
- using setsum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)]
- using setsum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
+ unfolding sum_clauses(2)[OF fin] and * using vu
+ using sum.cong [of s _ "\<lambda>x. u x *\<^sub>R x" "\<lambda>x. if x = a then v *\<^sub>R x else u x *\<^sub>R x", OF _ **(2)]
+ using sum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
apply auto
done
qed
@@ -1728,7 +1728,7 @@
have *:
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::real)"
"\<And>x y z. z = x - y \<longleftrightarrow> y + z = (x::'a)" by auto
- have "?lhs = {y. \<exists>u. setsum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
+ have "?lhs = {y. \<exists>u. sum u {a, b} = 1 \<and> (\<Sum>v\<in>{a, b}. u v *\<^sub>R v) = y}"
using affine_hull_finite[of "{a,b}"] by auto
also have "\<dots> = {y. \<exists>v u. u b = 1 - v \<and> u b *\<^sub>R b = y - v *\<^sub>R a}"
by (simp add: affine_hull_finite_step(2)[of "{b}" a])
@@ -1797,7 +1797,7 @@
apply (erule conjE)+
proof -
fix x t u
- assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
+ assume as: "finite t" "t \<noteq> {}" "t \<subseteq> insert a s" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = x"
have "(\<lambda>x. x - a) ` (t - {a}) \<subseteq> {x - a |x. x \<in> s}"
using as(3) by auto
then show "\<exists>v. x = a + v \<and> (\<exists>S u. finite S \<and> S \<subseteq> {x - a |x. x \<in> s} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = v)"
@@ -1808,8 +1808,8 @@
apply (rule conjI) using as(1) apply simp
apply (erule conjI)
using as(1)
- apply (simp add: setsum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
- setsum_subtractf scaleR_left.setsum[symmetric] setsum_diff1 scaleR_left_diff_distrib)
+ apply (simp add: sum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
+ sum_subtractf scaleR_left.sum[symmetric] sum_diff1 scaleR_left_diff_distrib)
unfolding as
apply simp
done
@@ -1828,16 +1828,16 @@
unfolding span_explicit by auto
define f where "f = (\<lambda>x. x + a) ` t"
have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
- unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def])
+ unfolding f_def using obt by (auto simp add: sum.reindex[unfolded inj_on_def])
have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
using f(2) assms by auto
- show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
+ show "\<exists>sa u. finite sa \<and> sa \<noteq> {} \<and> sa \<subseteq> insert a s \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
apply (rule_tac x = "insert a f" in exI)
- apply (rule_tac x = "\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
+ apply (rule_tac x = "\<lambda>x. if x=a then 1 - sum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
using assms and f
- unfolding setsum_clauses(2)[OF f(1)] and if_smult
- unfolding setsum.If_cases[OF f(1), of "\<lambda>x. x = a"]
- apply (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *)
+ unfolding sum_clauses(2)[OF f(1)] and if_smult
+ unfolding sum.If_cases[OF f(1), of "\<lambda>x. x = a"]
+ apply (auto simp add: sum_subtractf scaleR_left.sum algebra_simps *)
done
qed
@@ -2281,8 +2281,8 @@
lemma affine_dependent_explicit:
"affine_dependent p \<longleftrightarrow>
- (\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and>
- (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
+ (\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and>
+ (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)"
unfolding affine_dependent_def affine_hull_explicit mem_Collect_eq
apply rule
apply (erule bexE, erule exE, erule exE)
@@ -2293,25 +2293,25 @@
apply (erule bexE)
proof -
fix x s u
- assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ assume as: "x \<in> p" "finite s" "s \<noteq> {}" "s \<subseteq> p - {x}" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
have "x \<notin> s" using as(1,4) by auto
- show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+ show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
apply (rule_tac x="insert x s" in exI, rule_tac x="\<lambda>v. if v = x then - 1 else u v" in exI)
- unfolding if_smult and setsum_clauses(2)[OF as(2)] and setsum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
+ unfolding if_smult and sum_clauses(2)[OF as(2)] and sum_delta_notmem[OF \<open>x\<notin>s\<close>] and as
using as
apply auto
done
next
fix s u v
- assume as: "finite s" "s \<subseteq> p" "setsum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
+ assume as: "finite s" "s \<subseteq> p" "sum u s = 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0" "v \<in> s" "u v \<noteq> 0"
have "s \<noteq> {v}"
using as(3,6) by auto
- then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ then show "\<exists>x\<in>p. \<exists>s u. finite s \<and> s \<noteq> {} \<and> s \<subseteq> p - {x} \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
apply (rule_tac x=v in bexI)
apply (rule_tac x="s - {v}" in exI)
apply (rule_tac x="\<lambda>x. - (1 / u v) * u x" in exI)
- unfolding scaleR_scaleR[symmetric] and scaleR_right.setsum [symmetric]
- unfolding setsum_distrib_left[symmetric] and setsum_diff1[OF as(1)]
+ unfolding scaleR_scaleR[symmetric] and scaleR_right.sum [symmetric]
+ unfolding sum_distrib_left[symmetric] and sum_diff1[OF as(1)]
using as
apply auto
done
@@ -2321,24 +2321,24 @@
fixes s :: "'a::real_vector set"
assumes "finite s"
shows "affine_dependent s \<longleftrightarrow>
- (\<exists>u. setsum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = 0)"
+ (\<exists>u. sum u s = 0 \<and> (\<exists>v\<in>s. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) s = 0)"
(is "?lhs = ?rhs")
proof
have *: "\<And>vt u v. (if vt then u v else 0) *\<^sub>R v = (if vt then (u v) *\<^sub>R v else 0::'a)"
by auto
assume ?lhs
then obtain t u v where
- "finite t" "t \<subseteq> s" "setsum u t = 0" "v\<in>t" "u v \<noteq> 0" "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
+ "finite t" "t \<subseteq> s" "sum u t = 0" "v\<in>t" "u v \<noteq> 0" "(\<Sum>v\<in>t. u v *\<^sub>R v) = 0"
unfolding affine_dependent_explicit by auto
then show ?rhs
apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
- apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric]
+ apply auto unfolding * and sum.inter_restrict[OF assms, symmetric]
unfolding Int_absorb1[OF \<open>t\<subseteq>s\<close>]
apply auto
done
next
assume ?rhs
- then obtain u v where "setsum u s = 0" "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+ then obtain u v where "sum u s = 0" "v\<in>s" "u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
by auto
then show ?lhs unfolding affine_dependent_explicit
using assms by auto
@@ -2718,7 +2718,7 @@
shows "convex hull s =
{y. \<exists>k u x.
(\<forall>i\<in>{1::nat .. k}. 0 \<le> u i \<and> x i \<in> s) \<and>
- (setsum u {1..k} = 1) \<and> (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
+ (sum u {1..k} = 1) \<and> (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} = y)}"
(is "?xyz = ?hull")
apply (rule hull_unique)
apply rule
@@ -2743,7 +2743,7 @@
fix x k u y
assume assm:
"\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> s"
- "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
+ "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
show "x\<in>t"
unfolding assm(3) [symmetric]
apply (rule as(2)[unfolded convex, rule_format])
@@ -2755,10 +2755,10 @@
assume uv: "0 \<le> u" "0 \<le> v" "u + v = (1::real)"
assume xy: "x \<in> ?hull" "y \<in> ?hull"
from xy obtain k1 u1 x1 where
- x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "setsum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
+ x: "\<forall>i\<in>{1::nat..k1}. 0\<le>u1 i \<and> x1 i \<in> s" "sum u1 {Suc 0..k1} = 1" "(\<Sum>i = Suc 0..k1. u1 i *\<^sub>R x1 i) = x"
by auto
from xy obtain k2 u2 x2 where
- y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "setsum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
+ y: "\<forall>i\<in>{1::nat..k2}. 0\<le>u2 i \<and> x2 i \<in> s" "sum u2 {Suc 0..k2} = 1" "(\<Sum>i = Suc 0..k2. u2 i *\<^sub>R x2 i) = y"
by auto
have *: "\<And>P (x1::'a) x2 s1 s2 i.
(if P i then s1 else s2) *\<^sub>R (if P i then x1 else x2) = (if P i then s1 *\<^sub>R x1 else s2 *\<^sub>R x2)"
@@ -2779,9 +2779,9 @@
apply (rule, rule)
defer
apply rule
- unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
- setsum.reindex[OF inj] and o_def Collect_mem_eq
- unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_distrib_left[symmetric]
+ unfolding * and sum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
+ sum.reindex[OF inj] and o_def Collect_mem_eq
+ unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] sum_distrib_left[symmetric]
proof -
fix i
assume i: "i \<in> {1..k1+k2}"
@@ -2807,22 +2807,22 @@
fixes s :: "'a::real_vector set"
assumes "finite s"
shows "convex hull s = {y. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and>
- setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
+ sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
(is "?HULL = ?set")
proof (rule hull_unique, auto simp add: convex_def[of ?set])
fix x
assume "x \<in> s"
- then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x"
+ then show "\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = x"
apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI)
apply auto
- unfolding setsum.delta'[OF assms] and setsum_delta''[OF assms]
+ unfolding sum.delta'[OF assms] and sum_delta''[OF assms]
apply auto
done
next
fix u v :: real
assume uv: "0 \<le> u" "0 \<le> v" "u + v = 1"
- fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "setsum ux s = (1::real)"
- fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "setsum uy s = (1::real)"
+ fix ux assume ux: "\<forall>x\<in>s. 0 \<le> ux x" "sum ux s = (1::real)"
+ fix uy assume uy: "\<forall>x\<in>s. 0 \<le> uy x" "sum uy s = (1::real)"
{
fix x
assume "x\<in>s"
@@ -2832,15 +2832,15 @@
}
moreover
have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
- unfolding setsum.distrib and setsum_distrib_left[symmetric] and ux(2) uy(2)
+ unfolding sum.distrib and sum_distrib_left[symmetric] and ux(2) uy(2)
using uv(3) by auto
moreover
have "(\<Sum>x\<in>s. (u * ux x + v * uy x) *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
- unfolding scaleR_left_distrib and setsum.distrib and scaleR_scaleR[symmetric]
- and scaleR_right.setsum [symmetric]
+ unfolding scaleR_left_distrib and sum.distrib and scaleR_scaleR[symmetric]
+ and scaleR_right.sum [symmetric]
by auto
ultimately
- show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> setsum uc s = 1 \<and>
+ show "\<exists>uc. (\<forall>x\<in>s. 0 \<le> uc x) \<and> sum uc s = 1 \<and>
(\<Sum>x\<in>s. uc x *\<^sub>R x) = u *\<^sub>R (\<Sum>x\<in>s. ux x *\<^sub>R x) + v *\<^sub>R (\<Sum>x\<in>s. uy x *\<^sub>R x)"
apply (rule_tac x="\<lambda>x. u * ux x + v * uy x" in exI)
apply auto
@@ -2849,7 +2849,7 @@
fix t
assume t: "s \<subseteq> t" "convex t"
fix u
- assume u: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = (1::real)"
+ assume u: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = (1::real)"
then show "(\<Sum>x\<in>s. u x *\<^sub>R x) \<in> t"
using t(2)[unfolded convex_explicit, THEN spec[where x=s], THEN spec[where x=u]]
using assms and t(1) by auto
@@ -2861,14 +2861,14 @@
lemma convex_hull_explicit:
fixes p :: "'a::real_vector set"
shows "convex hull p =
- {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+ {y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
(is "?lhs = ?rhs")
proof -
{
fix x
assume "x\<in>?lhs"
then obtain k u y where
- obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
+ obt: "\<forall>i\<in>{1::nat..k}. 0 \<le> u i \<and> y i \<in> p" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R y i) = x"
unfolding convex_hull_indexed by auto
have fin: "finite {1..k}" by auto
@@ -2876,24 +2876,24 @@
{
fix j
assume "j\<in>{1..k}"
- then have "y j \<in> p" "0 \<le> setsum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
+ then have "y j \<in> p" "0 \<le> sum u {i. Suc 0 \<le> i \<and> i \<le> k \<and> y i = y j}"
using obt(1)[THEN bspec[where x=j]] and obt(2)
apply simp
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
using obt(1)
apply auto
done
}
moreover
- have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v}) = 1"
- unfolding setsum_image_gen[OF fin, symmetric] using obt(2) by auto
- moreover have "(\<Sum>v\<in>y ` {1..k}. setsum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
- using setsum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
- unfolding scaleR_left.setsum using obt(3) by auto
+ have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v}) = 1"
+ unfolding sum_image_gen[OF fin, symmetric] using obt(2) by auto
+ moreover have "(\<Sum>v\<in>y ` {1..k}. sum u {i \<in> {1..k}. y i = v} *\<^sub>R v) = x"
+ using sum_image_gen[OF fin, of "\<lambda>i. u i *\<^sub>R y i" y, symmetric]
+ unfolding scaleR_left.sum using obt(3) by auto
ultimately
- have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ have "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
apply (rule_tac x="y ` {1..k}" in exI)
- apply (rule_tac x="\<lambda>v. setsum u {i\<in>{1..k}. y i = v}" in exI)
+ apply (rule_tac x="\<lambda>v. sum u {i\<in>{1..k}. y i = v}" in exI)
apply auto
done
then have "x\<in>?rhs" by auto
@@ -2903,7 +2903,7 @@
fix y
assume "y\<in>?rhs"
then obtain s u where
- obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+ obt: "finite s" "s \<subseteq> p" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y"
by auto
obtain f where f: "inj_on f {1..card s}" "f ` {1..card s} = s"
@@ -2934,18 +2934,18 @@
then have "card {x. Suc 0 \<le> x \<and> x \<le> card s \<and> f x = y} = 1" by auto
then have "(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x)) = u y"
"(\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x) = u y *\<^sub>R y"
- by (auto simp add: setsum_constant_scaleR)
+ by (auto simp add: sum_constant_scaleR)
}
then have "(\<Sum>x = 1..card s. u (f x)) = 1" "(\<Sum>i = 1..card s. u (f i) *\<^sub>R f i) = y"
- unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
- and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
+ unfolding sum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
+ and sum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
unfolding f
- using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
- using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
+ using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x) *\<^sub>R f x)" "\<lambda>v. u v *\<^sub>R v"]
+ using sum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
unfolding obt(4,5)
by auto
ultimately
- have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> setsum u {1..k} = 1 \<and>
+ have "\<exists>k u x. (\<forall>i\<in>{1..k}. 0 \<le> u i \<and> x i \<in> p) \<and> sum u {1..k} = 1 \<and>
(\<Sum>i::nat = 1..k. u i *\<^sub>R x i) = y"
apply (rule_tac x="card s" in exI)
apply (rule_tac x="u \<circ> f" in exI)
@@ -2966,8 +2966,8 @@
fixes s :: "'a::real_vector set"
assumes "finite s"
shows
- "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> setsum u (insert a s) = w \<and> setsum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
- \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = w - v \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)"
+ "(\<exists>u. (\<forall>x\<in>insert a s. 0 \<le> u x) \<and> sum u (insert a s) = w \<and> sum (\<lambda>x. u x *\<^sub>R x) (insert a s) = y)
+ \<longleftrightarrow> (\<exists>v\<ge>0. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = w - v \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y - v *\<^sub>R a)"
(is "?lhs = ?rhs")
proof (rule, case_tac[!] "a\<in>s")
assume "a \<in> s"
@@ -2981,7 +2981,7 @@
next
assume ?lhs
then obtain u where
- u: "\<forall>x\<in>insert a s. 0 \<le> u x" "setsum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
+ u: "\<forall>x\<in>insert a s. 0 \<le> u x" "sum u (insert a s) = w" "(\<Sum>x\<in>insert a s. u x *\<^sub>R x) = y"
by auto
assume "a \<notin> s"
then show ?rhs
@@ -2989,7 +2989,7 @@
using u(1)[THEN bspec[where x=a]]
apply simp
apply (rule_tac x=u in exI)
- using u[unfolded setsum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
+ using u[unfolded sum_clauses(2)[OF assms]] and \<open>a\<notin>s\<close>
apply auto
done
next
@@ -2997,34 +2997,34 @@
then have *: "insert a s = s" by auto
have fin: "finite (insert a s)" using assms by auto
assume ?rhs
- then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+ then obtain v u where uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
by auto
show ?lhs
apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
- unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin]
- unfolding setsum_clauses(2)[OF assms]
+ unfolding scaleR_left_distrib and sum.distrib and sum_delta''[OF fin] and sum.delta'[OF fin]
+ unfolding sum_clauses(2)[OF assms]
using uv and uv(2)[THEN bspec[where x=a]] and \<open>a\<in>s\<close>
apply auto
done
next
assume ?rhs
then obtain v u where
- uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
+ uv: "v\<ge>0" "\<forall>x\<in>s. 0 \<le> u x" "sum u s = w - v" "(\<Sum>x\<in>s. u x *\<^sub>R x) = y - v *\<^sub>R a"
by auto
moreover
assume "a \<notin> s"
moreover
- have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s"
+ have "(\<Sum>x\<in>s. if a = x then v else u x) = sum u s"
and "(\<Sum>x\<in>s. (if a = x then v else u x) *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
- apply (rule_tac setsum.cong) apply rule
+ apply (rule_tac sum.cong) apply rule
defer
- apply (rule_tac setsum.cong) apply rule
+ apply (rule_tac sum.cong) apply rule
using \<open>a \<notin> s\<close>
apply auto
done
ultimately show ?lhs
apply (rule_tac x="\<lambda>x. if a = x then v else u x" in exI)
- unfolding setsum_clauses(2)[OF assms]
+ unfolding sum_clauses(2)[OF assms]
apply auto
done
qed
@@ -3150,12 +3150,12 @@
then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
by auto
moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
apply auto
done
have "(\<Sum>x\<in>insert a t. if x = a then - (\<Sum>x\<in>t. u (x - a)) else u (x - a)) = 0"
- unfolding setsum_clauses(2)[OF fin]
+ unfolding sum_clauses(2)[OF fin]
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
apply auto
unfolding *
@@ -3168,17 +3168,17 @@
apply auto
done
moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x) *\<^sub>R x) = (\<Sum>x\<in>t. Q x *\<^sub>R x)"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
apply auto
done
have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
- unfolding scaleR_left.setsum
- unfolding t_def and setsum.reindex[OF inj] and o_def
+ unfolding scaleR_left.sum
+ unfolding t_def and sum.reindex[OF inj] and o_def
using obt(5)
- by (auto simp add: setsum.distrib scaleR_right_distrib)
+ by (auto simp add: sum.distrib scaleR_right_distrib)
then have "(\<Sum>v\<in>insert a t. (if v = a then - (\<Sum>x\<in>t. u (x - a)) else u (v - a)) *\<^sub>R v) = 0"
- unfolding setsum_clauses(2)[OF fin]
+ unfolding sum_clauses(2)[OF fin]
using \<open>a\<notin>s\<close> \<open>t\<subseteq>s\<close>
by (auto simp add: *)
ultimately show ?thesis
@@ -3337,10 +3337,10 @@
assumes "0 \<in> affine hull S"
shows "dependent S"
proof -
- obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+ obtain s u where s_u: "finite s \<and> s \<noteq> {} \<and> s \<subseteq> S \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
using assms affine_hull_explicit[of S] by auto
then have "\<exists>v\<in>s. u v \<noteq> 0"
- using setsum_not_0[of "u" "s"] by auto
+ using sum_not_0[of "u" "s"] by auto
then have "finite s \<and> s \<subseteq> S \<and> (\<exists>v\<in>s. u v \<noteq> 0 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = 0)"
using s_u by auto
then show ?thesis
@@ -4131,18 +4131,18 @@
{ fix y
assume yt: "y \<in> affine hull t" and yu: "y \<in> affine hull u"
then obtain a b
- where a1 [simp]: "setsum a t = 1" and [simp]: "setsum (\<lambda>v. a v *\<^sub>R v) t = y"
- and [simp]: "setsum b u = 1" "setsum (\<lambda>v. b v *\<^sub>R v) u = y"
+ where a1 [simp]: "sum a t = 1" and [simp]: "sum (\<lambda>v. a v *\<^sub>R v) t = y"
+ and [simp]: "sum b u = 1" "sum (\<lambda>v. b v *\<^sub>R v) u = y"
by (auto simp: affine_hull_finite \<open>finite t\<close> \<open>finite u\<close>)
define c where "c x = (if x \<in> t then a x else if x \<in> u then -(b x) else 0)" for x
have [simp]: "s \<inter> t = t" "s \<inter> - t \<inter> u = u" using assms by auto
- have "setsum c s = 0"
- by (simp add: c_def comm_monoid_add_class.setsum.If_cases \<open>finite s\<close> setsum_negf)
+ have "sum c s = 0"
+ by (simp add: c_def comm_monoid_add_class.sum.If_cases \<open>finite s\<close> sum_negf)
moreover have "~ (\<forall>v\<in>s. c v = 0)"
- by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def setsum_not_0 zero_neq_one)
+ by (metis (no_types) IntD1 \<open>s \<inter> t = t\<close> a1 c_def sum_not_0 zero_neq_one)
moreover have "(\<Sum>v\<in>s. c v *\<^sub>R v) = 0"
- by (simp add: c_def if_smult setsum_negf
- comm_monoid_add_class.setsum.If_cases \<open>finite s\<close>)
+ by (simp add: c_def if_smult sum_negf
+ comm_monoid_add_class.sum.If_cases \<open>finite s\<close>)
ultimately have False
using assms \<open>finite s\<close> by (auto simp: affine_dependent_explicit)
}
@@ -4220,13 +4220,13 @@
fixes p :: "('a::euclidean_space) set"
shows "convex hull p =
{y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
- (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+ (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
unfolding convex_hull_explicit set_eq_iff mem_Collect_eq
proof (intro allI iffI)
fix y
let ?P = "\<lambda>n. \<exists>s u. finite s \<and> card s = n \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and>
- setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
- assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+ sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+ assume "\<exists>s u. finite s \<and> s \<subseteq> p \<and> (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
then obtain N where "?P N" by auto
then have "\<exists>n\<le>N. (\<forall>k<n. \<not> ?P k) \<and> ?P n"
apply (rule_tac ex_least_nat_le)
@@ -4235,7 +4235,7 @@
then obtain n where "?P n" and smallest: "\<forall>k<n. \<not> ?P k"
by blast
then obtain s u where obt: "finite s" "card s = n" "s\<subseteq>p" "\<forall>x\<in>s. 0 \<le> u x"
- "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
+ "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = y" by auto
have "card s \<le> aff_dim p + 1"
proof (rule ccontr, simp only: not_le)
@@ -4243,19 +4243,19 @@
then have "affine_dependent s"
using affine_dependent_biggerset[OF obt(1)] independent_card_le_aff_dim not_less obt(3)
by blast
- then obtain w v where wv: "setsum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
+ then obtain w v where wv: "sum w s = 0" "v\<in>s" "w v \<noteq> 0" "(\<Sum>v\<in>s. w v *\<^sub>R v) = 0"
using affine_dependent_explicit_finite[OF obt(1)] by auto
define i where "i = (\<lambda>v. (u v) / (- w v)) ` {v\<in>s. w v < 0}"
define t where "t = Min i"
have "\<exists>x\<in>s. w x < 0"
proof (rule ccontr, simp add: not_less)
assume as:"\<forall>x\<in>s. 0 \<le> w x"
- then have "setsum w (s - {v}) \<ge> 0"
- apply (rule_tac setsum_nonneg)
+ then have "sum w (s - {v}) \<ge> 0"
+ apply (rule_tac sum_nonneg)
apply auto
done
- then have "setsum w s > 0"
- unfolding setsum.remove[OF obt(1) \<open>v\<in>s\<close>]
+ then have "sum w s > 0"
+ unfolding sum.remove[OF obt(1) \<open>v\<in>s\<close>]
using as[THEN bspec[where x=v]] \<open>v\<in>s\<close> \<open>w v \<noteq> 0\<close> by auto
then show False using wv(1) by auto
qed
@@ -4291,12 +4291,12 @@
obtain a where "a \<in> s" and "t = (\<lambda>v. (u v) / (- w v)) a" and "w a < 0"
using Min_in[OF _ \<open>i\<noteq>{}\<close>] and obt(1) unfolding i_def t_def by auto
then have a: "a \<in> s" "u a + t * w a = 0" by auto
- have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
- unfolding setsum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
+ have *: "\<And>f. sum f (s - {a}) = sum f s - ((f a)::'b::ab_group_add)"
+ unfolding sum.remove[OF obt(1) \<open>a\<in>s\<close>] by auto
have "(\<Sum>v\<in>s. u v + t * w v) = 1"
- unfolding setsum.distrib wv(1) setsum_distrib_left[symmetric] obt(5) by auto
+ unfolding sum.distrib wv(1) sum_distrib_left[symmetric] obt(5) by auto
moreover have "(\<Sum>v\<in>s. u v *\<^sub>R v + (t * w v) *\<^sub>R v) - (u a *\<^sub>R a + (t * w a) *\<^sub>R a) = y"
- unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
+ unfolding sum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] wv(4)
using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
ultimately have "?P (n - 1)"
apply (rule_tac x="(s - {a})" in exI)
@@ -4308,7 +4308,7 @@
using smallest[THEN spec[where x="n - 1"]] by auto
qed
then show "\<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> aff_dim p + 1 \<and>
- (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
+ (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = y"
using obt by auto
qed auto
@@ -4332,7 +4332,7 @@
fixes p :: "('a::euclidean_space) set"
shows "convex hull p =
{y. \<exists>s u. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and>
- (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>v. u v *\<^sub>R v) s = y}"
+ (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> sum (\<lambda>v. u v *\<^sub>R v) s = y}"
(is "?lhs = ?rhs")
proof (intro set_eqI iffI)
fix x
@@ -4356,7 +4356,7 @@
fix x
assume "x \<in> convex hull p"
then obtain s u where "finite s" "s \<subseteq> p" "card s \<le> DIM('a) + 1"
- "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
+ "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "(\<Sum>v\<in>s. u v *\<^sub>R v) = x"
unfolding convex_hull_caratheodory by auto
then show "\<exists>s. finite s \<and> s \<subseteq> p \<and> card s \<le> DIM('a) + 1 \<and> x \<in> convex hull s"
apply (rule_tac x=s in exI)
@@ -5071,8 +5071,8 @@
unfolding mem_Collect_eq ball_simps(8)
proof (rule, rule)
fix a
- assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
- then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "setsum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a"
+ assume "\<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = a"
+ then obtain t u where obt: "finite t" "t\<subseteq>s" "\<forall>x\<in>t. 0 \<le> u x" "sum u t = 1" "(\<Sum>v\<in>t. u v *\<^sub>R v) = a"
by auto
from assms[unfolded open_contains_cball] obtain b
@@ -5083,7 +5083,7 @@
define i where "i = b ` t"
show "\<exists>e > 0.
- cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
+ cball a e \<subseteq> {y. \<exists>sa u. finite sa \<and> sa \<subseteq> s \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y}"
apply (rule_tac x = "Min i" in exI)
unfolding subset_eq
apply rule
@@ -5125,12 +5125,12 @@
have *: "inj_on (\<lambda>v. v + (y - a)) t"
unfolding inj_on_def by auto
have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
- unfolding setsum.reindex[OF *] o_def using obt(4) by auto
+ unfolding sum.reindex[OF *] o_def using obt(4) by auto
moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
- unfolding setsum.reindex[OF *] o_def using obt(4,5)
- by (simp add: setsum.distrib setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib)
+ unfolding sum.reindex[OF *] o_def using obt(4,5)
+ by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib)
ultimately
- show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> setsum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
+ show "\<exists>sa u. finite sa \<and> (\<forall>x\<in>sa. x \<in> s) \<and> (\<forall>x\<in>sa. 0 \<le> u x) \<and> sum u sa = 1 \<and> (\<Sum>v\<in>sa. u v *\<^sub>R v) = y"
apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI)
apply (rule_tac x="\<lambda>v. u (v - (y - a))" in exI)
using obt(1, 3)
@@ -6167,44 +6167,44 @@
lemma radon_ex_lemma:
assumes "finite c" "affine_dependent c"
- shows "\<exists>u. setsum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> setsum (\<lambda>v. u v *\<^sub>R v) c = 0"
+ shows "\<exists>u. sum u c = 0 \<and> (\<exists>v\<in>c. u v \<noteq> 0) \<and> sum (\<lambda>v. u v *\<^sub>R v) c = 0"
proof -
from assms(2)[unfolded affine_dependent_explicit]
obtain s u where
- "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+ "finite s" "s \<subseteq> c" "sum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
by blast
then show ?thesis
apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI)
- unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms(1), symmetric]
+ unfolding if_smult scaleR_zero_left and sum.inter_restrict[OF assms(1), symmetric]
apply (auto simp add: Int_absorb1)
done
qed
lemma radon_s_lemma:
assumes "finite s"
- and "setsum f s = (0::real)"
- shows "setsum f {x\<in>s. 0 < f x} = - setsum f {x\<in>s. f x < 0}"
+ and "sum f s = (0::real)"
+ shows "sum f {x\<in>s. 0 < f x} = - sum f {x\<in>s. f x < 0}"
proof -
have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
by auto
show ?thesis
- unfolding add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)]
- and setsum.distrib[symmetric] and *
+ unfolding add_eq_0_iff[symmetric] and sum.inter_filter[OF assms(1)]
+ and sum.distrib[symmetric] and *
using assms(2)
by assumption
qed
lemma radon_v_lemma:
assumes "finite s"
- and "setsum f s = 0"
+ and "sum f s = 0"
and "\<forall>x. g x = (0::real) \<longrightarrow> f x = (0::'a::euclidean_space)"
- shows "(setsum f {x\<in>s. 0 < g x}) = - setsum f {x\<in>s. g x < 0}"
+ shows "(sum f {x\<in>s. 0 < g x}) = - sum f {x\<in>s. g x < 0}"
proof -
have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
using assms(3) by auto
show ?thesis
- unfolding eq_neg_iff_add_eq_0 and setsum.inter_filter[OF assms(1)]
- and setsum.distrib[symmetric] and *
+ unfolding eq_neg_iff_add_eq_0 and sum.inter_filter[OF assms(1)]
+ and sum.distrib[symmetric] and *
using assms(2)
apply assumption
done
@@ -6214,12 +6214,12 @@
assumes "finite c" "affine_dependent c"
shows "\<exists>m p. m \<inter> p = {} \<and> m \<union> p = c \<and> (convex hull m) \<inter> (convex hull p) \<noteq> {}"
proof -
- obtain u v where uv: "setsum u c = 0" "v\<in>c" "u v \<noteq> 0" "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0"
+ obtain u v where uv: "sum u c = 0" "v\<in>c" "u v \<noteq> 0" "(\<Sum>v\<in>c. u v *\<^sub>R v) = 0"
using radon_ex_lemma[OF assms] by auto
have fin: "finite {x \<in> c. 0 < u x}" "finite {x \<in> c. 0 > u x}"
using assms(1) by auto
- define z where "z = inverse (setsum u {x\<in>c. u x > 0}) *\<^sub>R setsum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
- have "setsum u {x \<in> c. 0 < u x} \<noteq> 0"
+ define z where "z = inverse (sum u {x\<in>c. u x > 0}) *\<^sub>R sum (\<lambda>x. u x *\<^sub>R x) {x\<in>c. u x > 0}"
+ have "sum u {x \<in> c. 0 < u x} \<noteq> 0"
proof (cases "u v \<ge> 0")
case False
then have "u v < 0" by auto
@@ -6227,36 +6227,36 @@
proof (cases "\<exists>w\<in>{x \<in> c. 0 < u x}. u w > 0")
case True
then show ?thesis
- using setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
+ using sum_nonneg_eq_0_iff[of _ u, OF fin(1)] by auto
next
case False
- then have "setsum u c \<le> setsum (\<lambda>x. if x=v then u v else 0) c"
- apply (rule_tac setsum_mono)
+ then have "sum u c \<le> sum (\<lambda>x. if x=v then u v else 0) c"
+ apply (rule_tac sum_mono)
apply auto
done
then show ?thesis
- unfolding setsum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
+ unfolding sum.delta[OF assms(1)] using uv(2) and \<open>u v < 0\<close> and uv(1) by auto
qed
- qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
-
- then have *: "setsum u {x\<in>c. u x > 0} > 0"
+ qed (insert sum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
+
+ then have *: "sum u {x\<in>c. u x > 0} > 0"
unfolding less_le
apply (rule_tac conjI)
- apply (rule_tac setsum_nonneg)
- apply auto
- done
- moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
+ apply (rule_tac sum_nonneg)
+ apply auto
+ done
+ moreover have "sum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = sum u c"
"(\<Sum>x\<in>{x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}. u x *\<^sub>R x) = (\<Sum>x\<in>c. u x *\<^sub>R x)"
using assms(1)
- apply (rule_tac[!] setsum.mono_neutral_left)
- apply auto
- done
- then have "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
+ apply (rule_tac[!] sum.mono_neutral_left)
+ apply auto
+ done
+ then have "sum u {x \<in> c. 0 < u x} = - sum u {x \<in> c. 0 > u x}"
"(\<Sum>x\<in>{x \<in> c. 0 < u x}. u x *\<^sub>R x) = - (\<Sum>x\<in>{x \<in> c. 0 > u x}. u x *\<^sub>R x)"
unfolding eq_neg_iff_add_eq_0
using uv(1,4)
- by (auto simp add: setsum.union_inter_neutral[OF fin, symmetric])
- moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x"
+ by (auto simp add: sum.union_inter_neutral[OF fin, symmetric])
+ moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * - u x"
apply rule
apply (rule mult_nonneg_nonneg)
using *
@@ -6265,11 +6265,11 @@
ultimately have "z \<in> convex hull {v \<in> c. u v \<le> 0}"
unfolding convex_hull_explicit mem_Collect_eq
apply (rule_tac x="{v \<in> c. u v < 0}" in exI)
- apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * - u y" in exI)
- using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
- apply (auto simp add: setsum_negf setsum_distrib_left[symmetric])
- done
- moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * u x"
+ apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * - u y" in exI)
+ using assms(1) unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def
+ apply (auto simp add: sum_negf sum_distrib_left[symmetric])
+ done
+ moreover have "\<forall>x\<in>{v \<in> c. 0 < u v}. 0 \<le> inverse (sum u {x \<in> c. 0 < u x}) * u x"
apply rule
apply (rule mult_nonneg_nonneg)
using *
@@ -6278,11 +6278,11 @@
then have "z \<in> convex hull {v \<in> c. u v > 0}"
unfolding convex_hull_explicit mem_Collect_eq
apply (rule_tac x="{v \<in> c. 0 < u v}" in exI)
- apply (rule_tac x="\<lambda>y. inverse (setsum u {x\<in>c. u x > 0}) * u y" in exI)
+ apply (rule_tac x="\<lambda>y. inverse (sum u {x\<in>c. u x > 0}) * u y" in exI)
using assms(1)
- unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] and z_def
+ unfolding scaleR_scaleR[symmetric] scaleR_right.sum [symmetric] and z_def
using *
- apply (auto simp add: setsum_negf setsum_distrib_left[symmetric])
+ apply (auto simp add: sum_negf sum_distrib_left[symmetric])
done
ultimately show ?thesis
apply (rule_tac x="{v\<in>c. u v \<le> 0}" in exI)
@@ -6297,7 +6297,7 @@
proof -
from assms[unfolded affine_dependent_explicit]
obtain s u where
- "finite s" "s \<subseteq> c" "setsum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
+ "finite s" "s \<subseteq> c" "sum u s = 0" "\<exists>v\<in>s. u v \<noteq> 0" "(\<Sum>v\<in>s. u v *\<^sub>R v) = 0"
by blast
then have *: "finite s" "affine_dependent s" and s: "s \<subseteq> c"
unfolding affine_dependent_explicit by auto
@@ -6460,10 +6460,10 @@
lemma convex_on:
assumes "convex s"
shows "convex_on s f \<longleftrightarrow>
- (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> setsum u {1..k} = 1 \<longrightarrow>
- f (setsum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> setsum (\<lambda>i. u i * f(x i)) {1..k})"
+ (\<forall>k u x. (\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> x i \<in> s) \<and> sum u {1..k} = 1 \<longrightarrow>
+ f (sum (\<lambda>i. u i *\<^sub>R x i) {1..k} ) \<le> sum (\<lambda>i. u i * f(x i)) {1..k})"
unfolding convex_epigraph_convex[OF assms] convex epigraph_def Ball_def mem_Collect_eq
- unfolding fst_setsum snd_setsum fst_scaleR snd_scaleR
+ unfolding fst_sum snd_sum fst_scaleR snd_scaleR
apply safe
apply (drule_tac x=k in spec)
apply (drule_tac x=u in spec)
@@ -6473,7 +6473,7 @@
apply simp
apply (rule_tac y="\<Sum>i = 1..k. u i * f (fst (x i))" in order_trans)
defer
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply (erule_tac x=i in allE)
unfolding real_scaleR_def
apply (rule mult_left_mono)
@@ -6692,11 +6692,11 @@
fix x
assume "x \<in> convex hull s"
then obtain k u v where
- obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "setsum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
+ obt: "\<forall>i\<in>{1..k::nat}. 0 \<le> u i \<and> v i \<in> s" "sum u {1..k} = 1" "(\<Sum>i = 1..k. u i *\<^sub>R v i) = x"
unfolding convex_hull_indexed mem_Collect_eq by auto
have "(\<Sum>i = 1..k. u i * f (v i)) \<le> b"
- using setsum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
- unfolding setsum_distrib_right[symmetric] obt(2) mult_1
+ using sum_mono[of "{1..k}" "\<lambda>i. u i * f (v i)" "\<lambda>i. u i * b"]
+ unfolding sum_distrib_right[symmetric] obt(2) mult_1
apply (drule_tac meta_mp)
apply (rule mult_left_mono)
using assms(2) obt(1)
@@ -6709,8 +6709,8 @@
by auto
qed
-lemma inner_setsum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
- by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
+lemma inner_sum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
+ by (simp add: inner_sum_left sum.If_cases inner_Basis)
lemma convex_set_plus:
assumes "convex s" and "convex t" shows "convex (s + t)"
@@ -6722,7 +6722,7 @@
finally show "convex (s + t)" .
qed
-lemma convex_set_setsum:
+lemma convex_set_sum:
assumes "\<And>i. i \<in> A \<Longrightarrow> convex (B i)"
shows "convex (\<Sum>i\<in>A. B i)"
proof (cases "finite A")
@@ -6730,11 +6730,11 @@
by induct (auto simp: convex_set_plus)
qed auto
-lemma finite_set_setsum:
+lemma finite_set_sum:
assumes "finite A" and "\<forall>i\<in>A. finite (B i)" shows "finite (\<Sum>i\<in>A. B i)"
using assms by (induct set: finite, simp, simp add: finite_set_plus)
-lemma set_setsum_eq:
+lemma set_sum_eq:
"finite A \<Longrightarrow> (\<Sum>i\<in>A. B i) = {\<Sum>i\<in>A. f i |f. \<forall>i\<in>A. f i \<in> B i}"
apply (induct set: finite)
apply simp
@@ -6743,24 +6743,24 @@
apply (rule_tac x="fun_upd f x a" in exI)
apply simp
apply (rule_tac f="\<lambda>x. a + x" in arg_cong)
- apply (rule setsum.cong [OF refl])
+ apply (rule sum.cong [OF refl])
apply clarsimp
apply fast
done
-lemma box_eq_set_setsum_Basis:
+lemma box_eq_set_sum_Basis:
shows "{x. \<forall>i\<in>Basis. x\<bullet>i \<in> B i} = (\<Sum>i\<in>Basis. image (\<lambda>x. x *\<^sub>R i) (B i))"
- apply (subst set_setsum_eq [OF finite_Basis])
+ apply (subst set_sum_eq [OF finite_Basis])
apply safe
apply (fast intro: euclidean_representation [symmetric])
- apply (subst inner_setsum_left)
+ apply (subst inner_sum_left)
apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i")
apply (drule (1) bspec)
apply clarsimp
- apply (frule setsum.remove [OF finite_Basis])
+ apply (frule sum.remove [OF finite_Basis])
apply (erule trans)
apply simp
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply clarsimp
apply (frule_tac x=i in bspec, assumption)
apply (drule_tac x=x in bspec, assumption)
@@ -6770,7 +6770,7 @@
apply simp
done
-lemma convex_hull_set_setsum:
+lemma convex_hull_set_sum:
"convex hull (\<Sum>i\<in>A. B i) = (\<Sum>i\<in>A. convex hull (B i))"
proof (cases "finite A")
assume "finite A" then show ?thesis
@@ -6796,19 +6796,19 @@
(is "?int = convex hull ?points")
proof -
have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
- by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
+ by (simp add: inner_sum_left sum.If_cases inner_Basis)
have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
by (auto simp: cbox_def)
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)"
- by (simp only: box_eq_set_setsum_Basis)
+ by (simp only: box_eq_set_sum_Basis)
also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` (convex hull {0, 1}))"
by (simp only: convex_hull_eq_real_cbox zero_le_one)
also have "\<dots> = (\<Sum>i\<in>Basis. convex hull ((\<lambda>x. x *\<^sub>R i) ` {0, 1}))"
by (simp only: convex_hull_linear_image linear_scaleR_left)
also have "\<dots> = convex hull (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` {0, 1})"
- by (simp only: convex_hull_set_setsum)
+ by (simp only: convex_hull_set_sum)
also have "\<dots> = convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}}"
- by (simp only: box_eq_set_setsum_Basis)
+ by (simp only: box_eq_set_sum_Basis)
also have "convex hull {x. \<forall>i\<in>Basis. x\<bullet>i \<in> {0, 1}} = convex hull ?points"
by simp
finally show ?thesis .
@@ -6939,7 +6939,7 @@
op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` cbox 0 One"
using assms
apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric])
-apply (simp add: min_def max_def algebra_simps setsum_subtractf euclidean_representation)
+apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation)
done
lemma closed_interval_as_convex_hull:
@@ -6953,7 +6953,7 @@
obtain s::"'a set" where "finite s" and eq: "cbox 0 One = convex hull s"
by (blast intro: unit_cube_convex_hull)
have lin: "linear (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k)"
- by (rule linear_compose_setsum) (auto simp: algebra_simps linearI)
+ by (rule linear_compose_sum) (auto simp: algebra_simps linearI)
have "finite (op + a ` (\<lambda>x. \<Sum>k\<in>Basis. ((b \<bullet> k - a \<bullet> k) * (x \<bullet> k)) *\<^sub>R k) ` s)"
by (rule finite_imageI \<open>finite s\<close>)+
then show ?thesis
@@ -7141,13 +7141,13 @@
proof
define c where "c = (\<Sum>i\<in>Basis. (\<lambda>a. a *\<^sub>R i) ` {x\<bullet>i - d, x\<bullet>i + d})"
show "finite c"
- unfolding c_def by (simp add: finite_set_setsum)
+ unfolding c_def by (simp add: finite_set_sum)
have 1: "convex hull c = {a. \<forall>i\<in>Basis. a \<bullet> i \<in> cbox (x \<bullet> i - d) (x \<bullet> i + d)}"
- unfolding box_eq_set_setsum_Basis
- unfolding c_def convex_hull_set_setsum
+ unfolding box_eq_set_sum_Basis
+ unfolding c_def convex_hull_set_sum
apply (subst convex_hull_linear_image [symmetric])
apply (simp add: linear_iff scaleR_add_left)
- apply (rule setsum.cong [OF refl])
+ apply (rule sum.cong [OF refl])
apply (rule image_cong [OF _ refl])
apply (rule convex_hull_eq_real_cbox)
apply (cut_tac \<open>0 < d\<close>, simp)
@@ -7168,10 +7168,10 @@
unfolding 2
apply clarsimp
apply (subst euclidean_dist_l2)
- apply (rule order_trans [OF setL2_le_setsum])
+ apply (rule order_trans [OF setL2_le_sum])
apply (rule zero_le_dist)
unfolding e'
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply simp
done
qed
@@ -8232,18 +8232,18 @@
assumes "finite s"
and "0 \<notin> s"
shows "convex hull (insert 0 s) =
- {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s \<le> 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y)}"
+ {y. (\<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s \<le> 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y)}"
unfolding convex_hull_finite[OF finite.insertI[OF assms(1)]]
apply (rule set_eqI, rule)
unfolding mem_Collect_eq
apply (erule_tac[!] exE)
apply (erule_tac[!] conjE)+
- unfolding setsum_clauses(2)[OF \<open>finite s\<close>]
+ unfolding sum_clauses(2)[OF \<open>finite s\<close>]
apply (rule_tac x=u in exI)
defer
- apply (rule_tac x="\<lambda>x. if x = 0 then 1 - setsum u s else u x" in exI)
+ apply (rule_tac x="\<lambda>x. if x = 0 then 1 - sum u s else u x" in exI)
using assms(2)
- unfolding if_smult and setsum_delta_notmem[OF assms(2)]
+ unfolding if_smult and sum_delta_notmem[OF assms(2)]
apply auto
done
@@ -8268,19 +8268,19 @@
proof -
fix x :: "'a::euclidean_space"
fix u
- assume as: "\<forall>x\<in>?D. 0 \<le> u x" "setsum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
+ assume as: "\<forall>x\<in>?D. 0 \<le> u x" "sum u ?D \<le> 1" "(\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
have *: "\<forall>i\<in>Basis. i:d \<longrightarrow> u i = x\<bullet>i"
and "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
using as(3)
unfolding substdbasis_expansion_unique[OF assms]
by auto
- then have **: "setsum u ?D = setsum (op \<bullet> x) ?D"
+ then have **: "sum u ?D = sum (op \<bullet> x) ?D"
apply -
- apply (rule setsum.cong)
+ apply (rule sum.cong)
using assms
apply auto
done
- have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1"
+ have "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (op \<bullet> x) ?D \<le> 1"
proof (rule,rule)
fix i :: 'a
assume i: "i \<in> Basis"
@@ -8293,12 +8293,12 @@
using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close>[rule_format, OF i] by auto
ultimately show "0 \<le> x\<bullet>i" by auto
qed (insert as(2)[unfolded **], auto)
- then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
+ then show "(\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (op \<bullet> x) ?D \<le> 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
using \<open>(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)\<close> by auto
next
fix x :: "'a::euclidean_space"
- assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "setsum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
- show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> setsum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
+ assume as: "\<forall>i\<in>Basis. 0 \<le> x \<bullet> i" "sum (op \<bullet> x) ?D \<le> 1" "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)"
+ show "\<exists>u. (\<forall>x\<in>?D. 0 \<le> u x) \<and> sum u ?D \<le> 1 \<and> (\<Sum>x\<in>?D. u x *\<^sub>R x) = x"
using as d
unfolding substdbasis_expansion_unique[OF assms]
apply (rule_tac x="inner x" in exI)
@@ -8309,12 +8309,12 @@
lemma std_simplex:
"convex hull (insert 0 Basis) =
- {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis \<le> 1}"
+ {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 \<le> x\<bullet>i) \<and> sum (\<lambda>i. x\<bullet>i) Basis \<le> 1}"
using substd_simplex[of Basis] by auto
lemma interior_std_simplex:
"interior (convex hull (insert 0 Basis)) =
- {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> setsum (\<lambda>i. x\<bullet>i) Basis < 1}"
+ {x::'a::euclidean_space. (\<forall>i\<in>Basis. 0 < x\<bullet>i) \<and> sum (\<lambda>i. x\<bullet>i) Basis < 1}"
apply (rule set_eqI)
unfolding mem_interior std_simplex
unfolding subset_eq mem_Collect_eq Ball_def mem_ball
@@ -8326,8 +8326,8 @@
proof -
fix x :: 'a
fix e
- assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> setsum (op \<bullet> xa) Basis \<le> 1"
- show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> setsum (op \<bullet> x) Basis < 1"
+ assume "e > 0" and as: "\<forall>xa. dist x xa < e \<longrightarrow> (\<forall>x\<in>Basis. 0 \<le> xa \<bullet> x) \<and> sum (op \<bullet> xa) Basis \<le> 1"
+ show "(\<forall>xa\<in>Basis. 0 < x \<bullet> xa) \<and> sum (op \<bullet> x) Basis < 1"
apply safe
proof -
fix i :: 'a
@@ -8343,15 +8343,15 @@
have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis)) \<bullet> i =
x\<bullet>i + (if i = (SOME i. i\<in>Basis) then e/2 else 0)"
by (auto simp: SOME_Basis inner_Basis inner_simps)
- then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
- setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
- apply (rule_tac setsum.cong)
+ then have *: "sum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
+ sum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
+ apply (rule_tac sum.cong)
apply auto
done
- have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
- unfolding * setsum.distrib
+ have "sum (op \<bullet> x) Basis < sum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
+ unfolding * sum.distrib
using \<open>e > 0\<close> DIM_positive[where 'a='a]
- apply (subst setsum.delta')
+ apply (subst sum.delta')
apply (auto simp: SOME_Basis)
done
also have "\<dots> \<le> 1"
@@ -8359,13 +8359,13 @@
apply (drule_tac as[rule_format])
apply auto
done
- finally show "setsum (op \<bullet> x) Basis < 1" by auto
+ finally show "sum (op \<bullet> x) Basis < 1" by auto
qed
next
fix x :: 'a
- assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "setsum (op \<bullet> x) Basis < 1"
+ assume as: "\<forall>i\<in>Basis. 0 < x \<bullet> i" "sum (op \<bullet> x) Basis < 1"
obtain a :: 'b where "a \<in> UNIV" using UNIV_witness ..
- let ?d = "(1 - setsum (op \<bullet> x) Basis) / real (DIM('a))"
+ let ?d = "(1 - sum (op \<bullet> x) Basis) / real (DIM('a))"
have "Min ((op \<bullet> x) ` Basis) > 0"
apply (rule Min_grI)
using as(1)
@@ -8373,7 +8373,7 @@
done
moreover have "?d > 0"
using as(2) by (auto simp: Suc_le_eq DIM_positive)
- ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
+ ultimately show "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> (\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> sum (op \<bullet> y) Basis \<le> 1"
apply (rule_tac x="min (Min ((op \<bullet> x) ` Basis)) D" for D in exI)
apply rule
defer
@@ -8381,8 +8381,8 @@
proof -
fix y
assume y: "dist x y < min (Min (op \<bullet> x ` Basis)) ?d"
- have "setsum (op \<bullet> y) Basis \<le> setsum (\<lambda>i. x\<bullet>i + ?d) Basis"
- proof (rule setsum_mono)
+ have "sum (op \<bullet> y) Basis \<le> sum (\<lambda>i. x\<bullet>i + ?d) Basis"
+ proof (rule sum_mono)
fix i :: 'a
assume i: "i \<in> Basis"
then have "\<bar>y\<bullet>i - x\<bullet>i\<bar> < ?d"
@@ -8395,9 +8395,9 @@
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
qed
also have "\<dots> \<le> 1"
- unfolding setsum.distrib setsum_constant
+ unfolding sum.distrib sum_constant
by (auto simp add: Suc_le_eq)
- finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
+ finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> sum (op \<bullet> y) Basis \<le> 1"
proof safe
fix i :: 'a
assume i: "i \<in> Basis"
@@ -8419,13 +8419,13 @@
"a \<in> interior(convex hull (insert 0 Basis))"
proof -
let ?D = "Basis :: 'a set"
- let ?a = "setsum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
+ let ?a = "sum (\<lambda>b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis"
{
fix i :: 'a
assume i: "i \<in> Basis"
have "?a \<bullet> i = inverse (2 * real DIM('a))"
- by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
- (simp_all add: setsum.If_cases i) }
+ by (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
+ (simp_all add: sum.If_cases i) }
note ** = this
show ?thesis
apply (rule that[of ?a])
@@ -8436,15 +8436,15 @@
show "0 < ?a \<bullet> i"
unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive)
next
- have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
- apply (rule setsum.cong)
+ have "sum (op \<bullet> ?a) ?D = sum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
+ apply (rule sum.cong)
apply rule
apply auto
done
also have "\<dots> < 1"
- unfolding setsum_constant divide_inverse[symmetric]
+ unfolding sum_constant divide_inverse[symmetric]
by (auto simp add: field_simps)
- finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
+ finally show "sum (op \<bullet> ?a) ?D < 1" by auto
qed
qed
@@ -8478,11 +8478,11 @@
"ball x e \<inter> {xa. (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0)} \<subseteq> convex hull (insert 0 ?p)"
using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto
then have as: "\<forall>xa. dist x xa < e \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> xa\<bullet>i = 0) \<longrightarrow>
- (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> setsum (op \<bullet> xa) d \<le> 1"
+ (\<forall>i\<in>d. 0 \<le> xa \<bullet> i) \<and> sum (op \<bullet> xa) d \<le> 1"
unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto
have x0: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
using x rel_interior_subset substd_simplex[OF assms] by auto
- have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
+ have "(\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
apply rule
apply rule
proof -
@@ -8509,22 +8509,22 @@
by auto
have "\<And>i. i \<in> Basis \<Longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = x\<bullet>i + (if i = a then e/2 else 0)"
using a d by (auto simp: inner_simps inner_Basis)
- then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
- setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
- using d by (intro setsum.cong) auto
+ then have *: "sum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
+ sum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
+ using d by (intro sum.cong) auto
have "a \<in> Basis"
using \<open>a \<in> d\<close> d by auto
then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
using x0 d \<open>a\<in>d\<close> by (auto simp add: inner_add_left inner_Basis)
- have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
- unfolding * setsum.distrib
+ have "sum (op \<bullet> x) d < sum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
+ unfolding * sum.distrib
using \<open>e > 0\<close> \<open>a \<in> d\<close>
using \<open>finite d\<close>
- by (auto simp add: setsum.delta')
+ by (auto simp add: sum.delta')
also have "\<dots> \<le> 1"
using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
by auto
- finally show "setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
+ finally show "sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0)"
using x0 by auto
qed
}
@@ -8543,7 +8543,7 @@
unfolding substd_simplex[OF assms] by fastforce
obtain a where a: "a \<in> d"
using \<open>d \<noteq> {}\<close> by auto
- let ?d = "(1 - setsum (op \<bullet> x) d) / real (card d)"
+ let ?d = "(1 - sum (op \<bullet> x) d) / real (card d)"
have "0 < card d" using \<open>d \<noteq> {}\<close> \<open>finite d\<close>
by (simp add: card_gt_0_iff)
have "Min ((op \<bullet> x) ` d) > 0"
@@ -8564,8 +8564,8 @@
fix y :: 'a
assume y: "dist x y < min (Min (op \<bullet> x ` d)) ?d"
assume y2: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> y\<bullet>i = 0"
- have "setsum (op \<bullet> y) d \<le> setsum (\<lambda>i. x\<bullet>i + ?d) d"
- proof (rule setsum_mono)
+ have "sum (op \<bullet> y) d \<le> sum (\<lambda>i. x\<bullet>i + ?d) d"
+ proof (rule sum_mono)
fix i
assume "i \<in> d"
with d have i: "i \<in> Basis"
@@ -8579,9 +8579,9 @@
then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
qed
also have "\<dots> \<le> 1"
- unfolding setsum.distrib setsum_constant using \<open>0 < card d\<close>
+ unfolding sum.distrib sum_constant using \<open>0 < card d\<close>
by auto
- finally show "setsum (op \<bullet> y) d \<le> 1" .
+ finally show "sum (op \<bullet> y) d \<le> 1" .
fix i :: 'a
assume i: "i \<in> Basis"
@@ -8600,7 +8600,7 @@
}
ultimately have
"\<And>x. x \<in> rel_interior (convex hull insert 0 d) \<longleftrightarrow>
- x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> setsum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
+ x \<in> {x. (\<forall>i\<in>d. 0 < x \<bullet> i) \<and> sum (op \<bullet> x) d < 1 \<and> (\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0)}"
by blast
then show ?thesis by (rule set_eqI)
qed
@@ -8613,7 +8613,7 @@
where "a \<in> rel_interior (convex hull (insert 0 d))"
proof -
let ?D = d
- let ?a = "setsum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
+ let ?a = "sum (\<lambda>b::'a::euclidean_space. inverse (2 * real (card d)) *\<^sub>R b) ?D"
have "finite d"
apply (rule finite_subset)
using assms(2)
@@ -8625,10 +8625,10 @@
fix i
assume "i \<in> d"
have "?a \<bullet> i = inverse (2 * real (card d))"
- apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
- unfolding inner_setsum_left
- apply (rule setsum.cong)
- using \<open>i \<in> d\<close> \<open>finite d\<close> setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
+ apply (rule trans[of _ "sum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
+ unfolding inner_sum_left
+ apply (rule sum.cong)
+ using \<open>i \<in> d\<close> \<open>finite d\<close> sum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
d1 assms(2)
by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
}
@@ -8645,17 +8645,17 @@
by auto
finally show "0 < ?a \<bullet> i" by auto
next
- have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
- by (rule setsum.cong) (rule refl, rule **)
+ have "sum (op \<bullet> ?a) ?D = sum (\<lambda>i. inverse (2 * real (card d))) ?D"
+ by (rule sum.cong) (rule refl, rule **)
also have "\<dots> < 1"
- unfolding setsum_constant divide_real_def[symmetric]
+ unfolding sum_constant divide_real_def[symmetric]
by (auto simp add: field_simps)
- finally show "setsum (op \<bullet> ?a) ?D < 1" by auto
+ finally show "sum (op \<bullet> ?a) ?D < 1" by auto
next
fix i
assume "i \<in> Basis" and "i \<notin> d"
have "?a \<in> span d"
- proof (rule span_setsum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d])
+ proof (rule span_sum[of d "(\<lambda>b. b /\<^sub>R (2 * real (card d)))" d])
{
fix x :: "'a::euclidean_space"
assume "x \<in> d"
@@ -10232,20 +10232,20 @@
assumes "finite I"
assumes "\<forall>i\<in>I. convex (S i) \<and> (S i) \<noteq> {}"
shows "convex hull (\<Union>(S ` I)) =
- {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}"
+ {sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)}"
(is "?lhs = ?rhs")
proof -
have "?lhs \<supseteq> ?rhs"
proof
fix x
assume "x : ?rhs"
- then obtain c s where *: "setsum (\<lambda>i. c i *\<^sub>R s i) I = x" "setsum c I = 1"
+ then obtain c s where *: "sum (\<lambda>i. c i *\<^sub>R s i) I = x" "sum c I = 1"
"(\<forall>i\<in>I. c i \<ge> 0) \<and> (\<forall>i\<in>I. s i \<in> S i)" by auto
then have "\<forall>i\<in>I. s i \<in> convex hull (\<Union>(S ` I))"
using hull_subset[of "\<Union>(S ` I)" convex] by auto
then show "x \<in> ?lhs"
unfolding *(1)[symmetric]
- apply (subst convex_setsum[of I "convex hull \<Union>(S ` I)" c s])
+ apply (subst convex_sum[of I "convex hull \<Union>(S ` I)" c s])
using * assms convex_convex_hull
apply auto
done
@@ -10265,14 +10265,14 @@
fix x
assume "x \<in> S i"
define c where "c j = (if j = i then 1::real else 0)" for j
- then have *: "setsum c I = 1"
- using \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. 1::real"]
+ then have *: "sum c I = 1"
+ using \<open>finite I\<close> \<open>i \<in> I\<close> sum.delta[of I i "\<lambda>j::'a. 1::real"]
by auto
define s where "s j = (if j = i then x else p j)" for j
then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)"
using c_def by (auto simp add: algebra_simps)
- then have "x = setsum (\<lambda>i. c i *\<^sub>R s i) I"
- using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> setsum.delta[of I i "\<lambda>j::'a. x"]
+ then have "x = sum (\<lambda>i. c i *\<^sub>R s i) I"
+ using s_def c_def \<open>finite I\<close> \<open>i \<in> I\<close> sum.delta[of I i "\<lambda>j::'a. x"]
by auto
then have "x \<in> ?rhs"
apply auto
@@ -10292,20 +10292,20 @@
fix x y
assume xy: "x \<in> ?rhs \<and> y \<in> ?rhs"
from xy obtain c s where
- xc: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)"
+ xc: "x = sum (\<lambda>i. c i *\<^sub>R s i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> S i)"
by auto
from xy obtain d t where
- yc: "y = setsum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> setsum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)"
+ yc: "y = sum (\<lambda>i. d i *\<^sub>R t i) I \<and> (\<forall>i\<in>I. d i \<ge> 0) \<and> sum d I = 1 \<and> (\<forall>i\<in>I. t i \<in> S i)"
by auto
define e where "e i = u * c i + v * d i" for i
have ge0: "\<forall>i\<in>I. e i \<ge> 0"
using e_def xc yc uv by simp
- have "setsum (\<lambda>i. u * c i) I = u * setsum c I"
- by (simp add: setsum_distrib_left)
- moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I"
- by (simp add: setsum_distrib_left)
- ultimately have sum1: "setsum e I = 1"
- using e_def xc yc uv by (simp add: setsum.distrib)
+ have "sum (\<lambda>i. u * c i) I = u * sum c I"
+ by (simp add: sum_distrib_left)
+ moreover have "sum (\<lambda>i. v * d i) I = v * sum d I"
+ by (simp add: sum_distrib_left)
+ ultimately have sum1: "sum e I = 1"
+ using e_def xc yc uv by (simp add: sum.distrib)
define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)"
for i
{
@@ -10348,11 +10348,11 @@
}
then have *: "\<forall>i\<in>I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i"
by auto
- have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I"
- using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum.distrib)
- also have "\<dots> = setsum (\<lambda>i. e i *\<^sub>R q i) I"
+ have "u *\<^sub>R x + v *\<^sub>R y = sum (\<lambda>i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I"
+ using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib)
+ also have "\<dots> = sum (\<lambda>i. e i *\<^sub>R q i) I"
using * by auto
- finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (e i) *\<^sub>R (q i)) I"
+ finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\<lambda>i. (e i) *\<^sub>R (q i)) I"
by auto
then have "u *\<^sub>R x + v *\<^sub>R y \<in> ?rhs"
using ge0 sum1 qs by auto
@@ -10380,13 +10380,13 @@
then have "convex hull (\<Union>(s ` I)) = convex hull (S \<union> T)"
by auto
moreover have "convex hull \<Union>(s ` I) =
- {\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}"
+ {\<Sum> i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)}"
apply (subst convex_hull_finite_union[of I s])
using assms s_def I_def
apply auto
done
moreover have
- "{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs"
+ "{\<Sum>i\<in>I. c i *\<^sub>R sa i | c sa. (\<forall>i\<in>I. 0 \<le> c i) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. sa i \<in> s i)} \<le> ?rhs"
using s_def I_def by auto
ultimately show "?lhs \<subseteq> ?rhs" by auto
{
@@ -10432,8 +10432,8 @@
lemma rel_interior_sum_gen:
fixes S :: "'a \<Rightarrow> 'n::euclidean_space set"
assumes "\<forall>i\<in>I. convex (S i)"
- shows "rel_interior (setsum S I) = setsum (\<lambda>i. rel_interior (S i)) I"
- apply (subst setsum_set_cond_linear[of convex])
+ shows "rel_interior (sum S I) = sum (\<lambda>i. rel_interior (S i)) I"
+ apply (subst sum_set_cond_linear[of convex])
using rel_interior_sum rel_interior_sing[of "0"] assms
apply (auto simp add: convex_set_plus)
done
@@ -10460,14 +10460,14 @@
assumes "finite I"
and "I \<noteq> {}"
assumes "\<forall>i\<in>I. convex (S i) \<and> cone (S i) \<and> S i \<noteq> {}"
- shows "convex hull (\<Union>(S ` I)) = setsum S I"
+ shows "convex hull (\<Union>(S ` I)) = sum S I"
(is "?lhs = ?rhs")
proof -
{
fix x
assume "x \<in> ?lhs"
then obtain c xs where
- x: "x = setsum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)"
+ x: "x = sum (\<lambda>i. c i *\<^sub>R xs i) I \<and> (\<forall>i\<in>I. c i \<ge> 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. xs i \<in> S i)"
using convex_hull_finite_union[of I S] assms by auto
define s where "s i = c i *\<^sub>R xs i" for i
{
@@ -10477,24 +10477,24 @@
using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto
}
then have "\<forall>i\<in>I. s i \<in> S i" by auto
- moreover have "x = setsum s I" using x s_def by auto
+ moreover have "x = sum s I" using x s_def by auto
ultimately have "x \<in> ?rhs"
- using set_setsum_alt[of I S] assms by auto
+ using set_sum_alt[of I S] assms by auto
}
moreover
{
fix x
assume "x \<in> ?rhs"
- then obtain s where x: "x = setsum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
- using set_setsum_alt[of I S] assms by auto
+ then obtain s where x: "x = sum s I \<and> (\<forall>i\<in>I. s i \<in> S i)"
+ using set_sum_alt[of I S] assms by auto
define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i
- then have "x = setsum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I"
+ then have "x = sum (\<lambda>i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I"
using x assms by auto
moreover have "\<forall>i\<in>I. xs i \<in> S i"
using x xs_def assms by (simp add: cone_def)
moreover have "\<forall>i\<in>I. (1 :: real) / of_nat (card I) \<ge> 0"
by auto
- moreover have "setsum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1"
+ moreover have "sum (\<lambda>i. (1 :: real) / of_nat (card I)) I = 1"
using assms by auto
ultimately have "x \<in> ?lhs"
apply (subst convex_hull_finite_union[of I S])
@@ -10526,12 +10526,12 @@
using A_def I_def by auto
then have "convex hull (\<Union>(A ` I)) = convex hull (S \<union> T)"
by auto
- moreover have "convex hull \<Union>(A ` I) = setsum A I"
+ moreover have "convex hull \<Union>(A ` I) = sum A I"
apply (subst convex_hull_finite_union_cones[of I A])
using assms A_def I_def
apply auto
done
- moreover have "setsum A I = S + T"
+ moreover have "sum A I = S + T"
using A_def I_def
unfolding set_plus_def
apply auto
@@ -10546,7 +10546,7 @@
assumes "finite I"
and "\<forall>i\<in>I. convex (S i) \<and> S i \<noteq> {}"
shows "rel_interior (convex hull (\<Union>(S ` I))) =
- {setsum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and>
+ {sum (\<lambda>i. c i *\<^sub>R s i) I | c s. (\<forall>i\<in>I. c i > 0) \<and> sum c I = 1 \<and>
(\<forall>i\<in>I. s i \<in> rel_interior(S i))}"
(is "?lhs = ?rhs")
proof (cases "I = {}")
@@ -10621,7 +10621,7 @@
by blast
then have "K0 = convex hull (\<Union>(K ` I))"
using geq by auto
- also have "\<dots> = setsum K I"
+ also have "\<dots> = sum K I"
apply (subst convex_hull_finite_union_cones[of I K])
using assms
apply blast
@@ -10634,8 +10634,8 @@
using assms cone_cone_hull \<open>\<forall>i\<in>I. K i \<noteq> {}\<close> K_def
apply auto
done
- finally have "K0 = setsum K I" by auto
- then have *: "rel_interior K0 = setsum (\<lambda>i. (rel_interior (K i))) I"
+ finally have "K0 = sum K I" by auto
+ then have *: "rel_interior K0 = sum (\<lambda>i. (rel_interior (K i))) I"
using rel_interior_sum_gen[of I K] convK by auto
{
fix x
@@ -10643,8 +10643,8 @@
then have "(1::real, x) \<in> rel_interior K0"
using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull
by auto
- then obtain k where k: "(1::real, x) = setsum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
- using \<open>finite I\<close> * set_setsum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
+ then obtain k where k: "(1::real, x) = sum k I \<and> (\<forall>i\<in>I. k i \<in> rel_interior (K i))"
+ using \<open>finite I\<close> * set_sum_alt[of I "\<lambda>i. rel_interior (K i)"] by auto
{
fix i
assume "i \<in> I"
@@ -10656,8 +10656,8 @@
then obtain c s where
cs: "\<forall>i\<in>I. k i = (c i, c i *\<^sub>R s i) \<and> 0 < c i \<and> s i \<in> rel_interior (S i)"
by metis
- then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> setsum c I = 1"
- using k by (simp add: setsum_prod)
+ then have "x = (\<Sum>i\<in>I. c i *\<^sub>R s i) \<and> sum c I = 1"
+ using k by (simp add: sum_prod)
then have "x \<in> ?rhs"
using k
apply auto
@@ -10671,8 +10671,8 @@
{
fix x
assume "x \<in> ?rhs"
- then obtain c s where cs: "x = setsum (\<lambda>i. c i *\<^sub>R s i) I \<and>
- (\<forall>i\<in>I. c i > 0) \<and> setsum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))"
+ then obtain c s where cs: "x = sum (\<lambda>i. c i *\<^sub>R s i) I \<and>
+ (\<forall>i\<in>I. c i > 0) \<and> sum c I = 1 \<and> (\<forall>i\<in>I. s i \<in> rel_interior (S i))"
by auto
define k where "k i = (c i, c i *\<^sub>R s i)" for i
{
@@ -10682,10 +10682,10 @@
by auto
}
then have "(1::real, x) \<in> rel_interior K0"
- using K0_def * set_setsum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs
+ using K0_def * set_sum_alt[of I "(\<lambda>i. rel_interior (K i))"] assms k_def cs
apply auto
apply (rule_tac x = k in exI)
- apply (simp add: setsum_prod)
+ apply (simp add: sum_prod)
done
then have "x \<in> ?lhs"
using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x]
@@ -10801,18 +10801,18 @@
proof -
have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto
{ fix u v x
- assume uv: "setsum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "setsum v s = 1"
+ assume uv: "sum u t = 1" "\<forall>x\<in>s. 0 \<le> v x" "sum v s = 1"
"(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>v\<in>t. u v *\<^sub>R v)" "x \<in> t"
then have s: "s = (s - t) \<union> t" \<comment>\<open>split into separate cases\<close>
using assms by auto
have [simp]: "(\<Sum>x\<in>t. v x *\<^sub>R x) + (\<Sum>x\<in>s - t. v x *\<^sub>R x) = (\<Sum>x\<in>t. u x *\<^sub>R x)"
- "setsum v t + setsum v (s - t) = 1"
+ "sum v t + sum v (s - t) = 1"
using uv fin s
- by (auto simp: setsum.union_disjoint [symmetric] Un_commute)
+ by (auto simp: sum.union_disjoint [symmetric] Un_commute)
have "(\<Sum>x\<in>s. if x \<in> t then v x - u x else v x) = 0"
"(\<Sum>x\<in>s. (if x \<in> t then v x - u x else v x) *\<^sub>R x) = 0"
using uv fin
- by (subst s, subst setsum.union_disjoint, auto simp: algebra_simps setsum_subtractf)+
+ by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+
} note [simp] = this
have "convex hull t \<subseteq> affine hull t"
using convex_hull_subset_affine_hull by blast
@@ -10890,14 +10890,14 @@
lemma explicit_subset_rel_interior_convex_hull:
fixes s :: "'a::euclidean_space set"
shows "finite s
- \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
+ \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}
\<subseteq> rel_interior (convex hull s)"
by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
lemma explicit_subset_rel_interior_convex_hull_minimal:
fixes s :: "'a::euclidean_space set"
shows "finite s
- \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}
+ \<Longrightarrow> {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}
\<subseteq> rel_interior (convex hull s)"
by (force simp add: rel_interior_convex_hull_union [where S="\<lambda>x. {x}" and I=s, simplified])
@@ -10905,7 +10905,7 @@
fixes s :: "'a::euclidean_space set"
assumes "~ affine_dependent s"
shows "rel_interior(convex hull s) =
- {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
+ {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
(is "?lhs = ?rhs")
proof
show "?rhs \<le> ?lhs"
@@ -10926,14 +10926,14 @@
have "(\<Sum>x\<in>s. if x = a then - d else if x = b then d else 0) = 0"
"(\<Sum>x\<in>s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a"
using ab fs
- by (subst s, subst setsum.union_disjoint, auto)+
+ by (subst s, subst sum.union_disjoint, auto)+
} note [simp] = this
{ fix y
assume y: "y \<in> convex hull s" "y \<notin> ?rhs"
{ fix u T a
- assume ua: "\<forall>x\<in>s. 0 \<le> u x" "setsum u s = 1" "\<not> 0 < u a" "a \<in> s"
+ assume ua: "\<forall>x\<in>s. 0 \<le> u x" "sum u s = 1" "\<not> 0 < u a" "a \<in> s"
and yT: "y = (\<Sum>x\<in>s. u x *\<^sub>R x)" "y \<in> T" "open T"
- and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}"
+ and sb: "T \<inter> affine hull s \<subseteq> {w. \<exists>u. (\<forall>x\<in>s. 0 \<le> u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = w}"
have ua0: "u a = 0"
using ua by auto
obtain b where b: "b\<in>s" "a \<noteq> b"
@@ -10949,7 +10949,7 @@
then have "y - d *\<^sub>R (a - b) \<in> T \<inter> affine hull s"
using d e yT by auto
then obtain v where "\<forall>x\<in>s. 0 \<le> v x"
- "setsum v s = 1"
+ "sum v s = 1"
"(\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x) - d *\<^sub>R (a - b)"
using subsetD [OF sb] yT
by auto
@@ -10958,7 +10958,7 @@
apply (simp add: affine_dependent_explicit_finite fs)
apply (drule_tac x="\<lambda>x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec)
using ua b d
- apply (auto simp: algebra_simps setsum_subtractf setsum.distrib)
+ apply (auto simp: algebra_simps sum_subtractf sum.distrib)
done
} note * = this
have "y \<notin> rel_interior (convex hull s)"
@@ -10979,7 +10979,7 @@
"~ affine_dependent s
==> interior(convex hull s) =
(if card(s) \<le> DIM('a) then {}
- else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
+ else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify)
apply (rule trans [of _ "rel_interior(convex hull s)"])
apply (simp add: affine_hull_convex_hull affine_independent_span_gt rel_interior_interior)
@@ -10991,10 +10991,10 @@
shows
"interior(convex hull s) =
(if card(s) \<le> DIM('a) then {}
- else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> setsum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
+ else {y. \<exists>u. (\<forall>x \<in> s. 0 < u x \<and> u x < 1) \<and> sum u s = 1 \<and> (\<Sum>x\<in>s. u x *\<^sub>R x) = y})"
proof -
{ fix u :: "'a \<Rightarrow> real" and a
- assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "setsum u s = 1" and a: "a \<in> s"
+ assume "card Basis < card s" and u: "\<And>x. x\<in>s \<Longrightarrow> 0 < u x" "sum u s = 1" and a: "a \<in> s"
then have cs: "Suc 0 < card s"
by (metis DIM_positive less_trans_Suc)
obtain b where b: "b \<in> s" "a \<noteq> b"
@@ -11007,10 +11007,10 @@
then show thesis
by (blast intro: that)
qed
- have "u a + u b \<le> setsum u {a,b}"
+ have "u a + u b \<le> sum u {a,b}"
using a b by simp
- also have "... \<le> setsum u s"
- apply (rule Groups_Big.setsum_mono2)
+ also have "... \<le> sum u s"
+ apply (rule Groups_Big.sum_mono2)
using a b u
apply (auto simp: less_imp_le aff_independent_finite assms)
done
@@ -11156,7 +11156,7 @@
fixes s :: "'a::euclidean_space set"
assumes "~ affine_dependent s"
shows "rel_frontier(convex hull s) =
- {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
+ {y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (\<exists>x \<in> s. u x = 0) \<and> sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
proof -
have fs: "finite s"
using assms by (simp add: aff_independent_finite)
@@ -11172,7 +11172,7 @@
apply (rule_tac x=s in exI)
apply (auto simp: fs)
apply (rule_tac x = "\<lambda>x. u x - v x" in exI)
- apply (force simp: setsum_subtractf scaleR_diff_left)
+ apply (force simp: sum_subtractf scaleR_diff_left)
done
qed
@@ -11181,7 +11181,7 @@
assumes "~ affine_dependent s"
shows "frontier(convex hull s) =
{y. \<exists>u. (\<forall>x \<in> s. 0 \<le> u x) \<and> (DIM ('a) < card s \<longrightarrow> (\<exists>x \<in> s. u x = 0)) \<and>
- setsum u s = 1 \<and> setsum (\<lambda>x. u x *\<^sub>R x) s = y}"
+ sum u s = 1 \<and> sum (\<lambda>x. u x *\<^sub>R x) s = y}"
proof -
have fs: "finite s"
using assms by (simp add: aff_independent_finite)
@@ -11211,21 +11211,21 @@
have fs: "finite s"
using assms by (simp add: aff_independent_finite)
{ fix u a
- have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> setsum u s = 1 \<Longrightarrow>
+ have "\<forall>x\<in>s. 0 \<le> u x \<Longrightarrow> a \<in> s \<Longrightarrow> u a = 0 \<Longrightarrow> sum u s = 1 \<Longrightarrow>
\<exists>x v. x \<in> s \<and>
(\<forall>x\<in>s - {x}. 0 \<le> v x) \<and>
- setsum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
+ sum v (s - {x}) = 1 \<and> (\<Sum>x\<in>s - {x}. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)"
apply (rule_tac x=a in exI)
apply (rule_tac x=u in exI)
- apply (simp add: Groups_Big.setsum_diff1 fs)
+ apply (simp add: Groups_Big.sum_diff1 fs)
done }
moreover
{ fix a u
- have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> setsum u (s - {a}) = 1 \<Longrightarrow>
+ have "a \<in> s \<Longrightarrow> \<forall>x\<in>s - {a}. 0 \<le> u x \<Longrightarrow> sum u (s - {a}) = 1 \<Longrightarrow>
\<exists>v. (\<forall>x\<in>s. 0 \<le> v x) \<and>
- (\<exists>x\<in>s. v x = 0) \<and> setsum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)"
+ (\<exists>x\<in>s. v x = 0) \<and> sum v s = 1 \<and> (\<Sum>x\<in>s. v x *\<^sub>R x) = (\<Sum>x\<in>s - {a}. u x *\<^sub>R x)"
apply (rule_tac x="\<lambda>x. if x = a then 0 else u x" in exI)
- apply (auto simp: setsum.If_cases Diff_eq if_smult fs)
+ apply (auto simp: sum.If_cases Diff_eq if_smult fs)
done }
ultimately show ?thesis
using assms
@@ -12496,7 +12496,7 @@
have "x = (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b)"
by (simp add: euclidean_representation)
also have "... = (\<Sum>b \<in> Basis - {k}. (x \<bullet> b) *\<^sub>R b)"
- by (auto simp: setsum.remove [of _ k] inner_commute assms that)
+ by (auto simp: sum.remove [of _ k] inner_commute assms that)
finally have "x = (\<Sum>b\<in>Basis - {k}. (x \<bullet> b) *\<^sub>R b)" .
then show ?thesis
by (simp add: Linear_Algebra.span_finite) metis
@@ -12960,20 +12960,20 @@
proof -
have [simp]: "finite s" "finite t"
using aff_independent_finite assms by blast+
- have "setsum u (s \<inter> t) = 1 \<and>
+ have "sum u (s \<inter> t) = 1 \<and>
(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)"
- if [simp]: "setsum u s = 1"
- "setsum v t = 1"
+ if [simp]: "sum u s = 1"
+ "sum v t = 1"
and eq: "(\<Sum>x\<in>t. v x *\<^sub>R x) = (\<Sum>x\<in>s. u x *\<^sub>R x)" for u v
proof -
define f where "f x = (if x \<in> s then u x else 0) - (if x \<in> t then v x else 0)" for x
- have "setsum f (s \<union> t) = 0"
- apply (simp add: f_def setsum_Un setsum_subtractf)
- apply (simp add: setsum.inter_restrict [symmetric] Int_commute)
+ have "sum f (s \<union> t) = 0"
+ apply (simp add: f_def sum_Un sum_subtractf)
+ apply (simp add: sum.inter_restrict [symmetric] Int_commute)
done
moreover have "(\<Sum>x\<in>(s \<union> t). f x *\<^sub>R x) = 0"
- apply (simp add: f_def setsum_Un scaleR_left_diff_distrib setsum_subtractf)
- apply (simp add: if_smult setsum.inter_restrict [symmetric] Int_commute eq
+ apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf)
+ apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq
cong del: if_weak_cong)
done
ultimately have "\<And>v. v \<in> s \<union> t \<Longrightarrow> f v = 0"
@@ -12981,12 +12981,12 @@
by blast
then have u [simp]: "\<And>x. x \<in> s \<Longrightarrow> u x = (if x \<in> t then v x else 0)"
by (simp add: f_def) presburger
- have "setsum u (s \<inter> t) = setsum u s"
- by (simp add: setsum.inter_restrict)
- then have "setsum u (s \<inter> t) = 1"
+ have "sum u (s \<inter> t) = sum u s"
+ by (simp add: sum.inter_restrict)
+ then have "sum u (s \<inter> t) = 1"
using that by linarith
moreover have "(\<Sum>v\<in>s \<inter> t. u v *\<^sub>R v) = (\<Sum>v\<in>s. u v *\<^sub>R v)"
- by (auto simp: if_smult setsum.inter_restrict intro: setsum.cong)
+ by (auto simp: if_smult sum.inter_restrict intro: sum.cong)
ultimately show ?thesis
by force
qed
@@ -13063,19 +13063,19 @@
have [simp]: "finite S"
by (simp add: aff_independent_finite assms)
then obtain b where b0: "\<And>s. s \<in> S \<Longrightarrow> 0 \<le> b s"
- and b1: "setsum b S = 1" and aeq: "a = (\<Sum>s\<in>S. b s *\<^sub>R s)"
+ and b1: "sum b S = 1" and aeq: "a = (\<Sum>s\<in>S. b s *\<^sub>R s)"
using a by (auto simp: convex_hull_finite)
have fin [simp]: "finite T" "finite T'"
using assms infinite_super \<open>finite S\<close> by blast+
then obtain c c' where c0: "\<And>t. t \<in> insert a T \<Longrightarrow> 0 \<le> c t"
- and c1: "setsum c (insert a T) = 1"
+ and c1: "sum c (insert a T) = 1"
and xeq: "x = (\<Sum>t \<in> insert a T. c t *\<^sub>R t)"
and c'0: "\<And>t. t \<in> insert a T' \<Longrightarrow> 0 \<le> c' t"
- and c'1: "setsum c' (insert a T') = 1"
+ and c'1: "sum c' (insert a T') = 1"
and x'eq: "x = (\<Sum>t \<in> insert a T'. c' t *\<^sub>R t)"
using x x' by (auto simp: convex_hull_finite)
with fin anot
- have sumTT': "setsum c T = 1 - c a" "setsum c' T' = 1 - c' a"
+ have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a"
and wsumT: "(\<Sum>t \<in> T. c t *\<^sub>R t) = x - c a *\<^sub>R a"
by simp_all
have wsumT': "(\<Sum>t \<in> T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a"
@@ -13083,23 +13083,23 @@
define cc where "cc \<equiv> \<lambda>x. if x \<in> T then c x else 0"
define cc' where "cc' \<equiv> \<lambda>x. if x \<in> T' then c' x else 0"
define dd where "dd \<equiv> \<lambda>x. cc x - cc' x + (c a - c' a) * b x"
- have sumSS': "setsum cc S = 1 - c a" "setsum cc' S = 1 - c' a"
+ have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a"
unfolding cc_def cc'_def using S
- by (simp_all add: Int_absorb1 Int_absorb2 setsum_subtractf setsum.inter_restrict [symmetric] sumTT')
+ by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT')
have wsumSS: "(\<Sum>t \<in> S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\<Sum>t \<in> S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a"
unfolding cc_def cc'_def using S
- by (simp_all add: Int_absorb1 Int_absorb2 if_smult setsum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong)
- have sum_dd0: "setsum dd S = 0"
+ by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong)
+ have sum_dd0: "sum dd S = 0"
unfolding dd_def using S
- by (simp add: sumSS' comm_monoid_add_class.setsum.distrib setsum_subtractf
- algebra_simps setsum_distrib_right [symmetric] b1)
+ by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf
+ algebra_simps sum_distrib_right [symmetric] b1)
have "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\<Sum>v\<in>S. b v *\<^sub>R v)" for x
- by (simp add: pth_5 real_vector.scale_setsum_right mult.commute)
+ by (simp add: pth_5 real_vector.scale_sum_right mult.commute)
then have *: "(\<Sum>v\<in>S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x
using aeq by blast
have "(\<Sum>v \<in> S. dd v *\<^sub>R v) = 0"
unfolding dd_def using S
- by (simp add: * wsumSS setsum.distrib setsum_subtractf algebra_simps)
+ by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps)
then have dd0: "dd v = 0" if "v \<in> S" for v
using naff that \<open>finite S\<close> sum_dd0 unfolding affine_dependent_explicit
apply (simp only: not_ex)
@@ -13110,7 +13110,7 @@
then show ?thesis
proof cases
case 1
- then have "setsum cc S \<le> setsum cc' S"
+ then have "sum cc S \<le> sum cc' S"
by (simp add: sumSS')
then have le: "cc x \<le> cc' x" if "x \<in> S" for x
using dd0 [OF that] 1 b0 mult_left_mono that
@@ -13124,32 +13124,32 @@
by (simp add: c0 cc_def)
show "0 \<le> (cc(a := c a)) a"
by (simp add: c0)
- have "setsum (cc(a := c a)) (insert a (T \<inter> T')) = c a + setsum (cc(a := c a)) (T \<inter> T')"
+ have "sum (cc(a := c a)) (insert a (T \<inter> T')) = c a + sum (cc(a := c a)) (T \<inter> T')"
by (simp add: anot)
- also have "... = c a + setsum (cc(a := c a)) S"
+ also have "... = c a + sum (cc(a := c a)) S"
apply simp
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a + (1 - c a)"
- by (metis \<open>a \<notin> S\<close> fun_upd_other setsum.cong sumSS')
- finally show "setsum (cc(a := c a)) (insert a (T \<inter> T')) = 1"
+ by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS')
+ finally show "sum (cc(a := c a)) (insert a (T \<inter> T')) = 1"
by simp
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc(a := c a)) x *\<^sub>R x)"
by (simp add: anot)
also have "... = c a *\<^sub>R a + (\<Sum>x \<in> S. (cc(a := c a)) x *\<^sub>R x)"
apply simp
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a"
- by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult setsum_delta_notmem)
+ by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem)
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc(a := c a)) x *\<^sub>R x) = x"
by simp
qed
next
case 2
- then have "setsum cc' S \<le> setsum cc S"
+ then have "sum cc' S \<le> sum cc S"
by (simp add: sumSS')
then have le: "cc' x \<le> cc x" if "x \<in> S" for x
using dd0 [OF that] 2 b0 mult_left_mono that
@@ -13163,26 +13163,26 @@
by (simp add: c'0 cc'_def)
show "0 \<le> (cc'(a := c' a)) a"
by (simp add: c'0)
- have "setsum (cc'(a := c' a)) (insert a (T \<inter> T')) = c' a + setsum (cc'(a := c' a)) (T \<inter> T')"
+ have "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = c' a + sum (cc'(a := c' a)) (T \<inter> T')"
by (simp add: anot)
- also have "... = c' a + setsum (cc'(a := c' a)) S"
+ also have "... = c' a + sum (cc'(a := c' a)) S"
apply simp
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c' a + (1 - c' a)"
- by (metis \<open>a \<notin> S\<close> fun_upd_other setsum.cong sumSS')
- finally show "setsum (cc'(a := c' a)) (insert a (T \<inter> T')) = 1"
+ by (metis \<open>a \<notin> S\<close> fun_upd_other sum.cong sumSS')
+ finally show "sum (cc'(a := c' a)) (insert a (T \<inter> T')) = 1"
by simp
have "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\<Sum>x \<in> T \<inter> T'. (cc'(a := c' a)) x *\<^sub>R x)"
by (simp add: anot)
also have "... = c' a *\<^sub>R a + (\<Sum>x \<in> S. (cc'(a := c' a)) x *\<^sub>R x)"
apply simp
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
using \<open>T \<subseteq> S\<close> apply (auto simp: \<open>a \<notin> S\<close> cc0)
done
also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a"
- by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult setsum_delta_notmem)
+ by (simp add: wsumSS \<open>a \<notin> S\<close> if_smult sum_delta_notmem)
finally show "(\<Sum>x\<in>insert a (T \<inter> T'). (cc'(a := c' a)) x *\<^sub>R x) = x"
by simp
qed
@@ -13458,10 +13458,10 @@
have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> S"
apply (subst eq)
apply (rule subspace_neg [OF \<open>subspace S\<close>])
- apply (rule subspace_setsum [OF \<open>subspace S\<close>])
+ apply (rule subspace_sum [OF \<open>subspace S\<close>])
by (meson subsetCE subspace_mul \<open>C \<subseteq> S\<close> \<open>subspace S\<close>)
moreover have "(\<Sum>v\<in>D - C. a v *\<^sub>R v) \<in> T"
- apply (rule subspace_setsum [OF \<open>subspace T\<close>])
+ apply (rule subspace_sum [OF \<open>subspace T\<close>])
by (meson DiffD1 \<open>D \<subseteq> T\<close> \<open>subspace T\<close> subset_eq subspace_def)
ultimately have "(\<Sum>v \<in> D-C. a v *\<^sub>R v) \<in> span B"
using B by blast
@@ -13475,12 +13475,12 @@
have f2: "(\<Sum>v\<in>C \<inter> D. e v *\<^sub>R v) = (\<Sum>v\<in>D - C. a v *\<^sub>R v)"
using Beq e by presburger
have f3: "(\<Sum>v\<in>C \<union> D. a v *\<^sub>R v) = (\<Sum>v\<in>C - D. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v) + (\<Sum>v\<in>C \<inter> D. a v *\<^sub>R v)"
- using \<open>finite C\<close> \<open>finite D\<close> setsum.union_diff2 by blast
+ using \<open>finite C\<close> \<open>finite D\<close> sum.union_diff2 by blast
have f4: "(\<Sum>v\<in>C \<union> (D - C). a v *\<^sub>R v) = (\<Sum>v\<in>C. a v *\<^sub>R v) + (\<Sum>v\<in>D - C. a v *\<^sub>R v)"
- by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff setsum.union_disjoint)
+ by (meson Diff_disjoint \<open>finite C\<close> \<open>finite D\<close> finite_Diff sum.union_disjoint)
have "(\<Sum>v\<in>C. cc v *\<^sub>R v) = 0"
using 0 f2 f3 f4
- apply (simp add: cc_def Beq if_smult \<open>finite C\<close> setsum.If_cases algebra_simps setsum.distrib)
+ apply (simp add: cc_def Beq if_smult \<open>finite C\<close> sum.If_cases algebra_simps sum.distrib)
apply (simp add: add.commute add.left_commute diff_eq)
done
then have "\<And>v. v \<in> C \<Longrightarrow> cc v = 0"
@@ -13498,14 +13498,14 @@
qed
with 0 have Dcc0: "(\<Sum>v\<in>D. a v *\<^sub>R v) = 0"
apply (subst Deq)
- by (simp add: \<open>finite B\<close> \<open>finite D\<close> setsum_Un)
+ by (simp add: \<open>finite B\<close> \<open>finite D\<close> sum_Un)
then have D0: "\<And>v. v \<in> D \<Longrightarrow> a v = 0"
using independent_explicit \<open>independent D\<close> by blast
show ?thesis
using v C0 D0 Beq by blast
qed
then have "independent (C \<union> (D - C))"
- by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> setsum_Un del: Un_Diff_cancel)
+ by (simp add: independent_explicit \<open>finite C\<close> \<open>finite D\<close> sum_Un del: Un_Diff_cancel)
then have indCUD: "independent (C \<union> D)" by simp
have "dim (S \<inter> T) = card B"
by (rule dim_unique [OF B indB refl])
@@ -13660,8 +13660,8 @@
then have "a \<bullet> a = a \<bullet> (\<Sum>v\<in>T. U v *\<^sub>R v)"
by simp
also have "... = 0"
- apply (simp add: inner_setsum_right)
- apply (rule comm_monoid_add_class.setsum.neutral)
+ apply (simp add: inner_sum_right)
+ apply (rule comm_monoid_add_class.sum.neutral)
by (metis "0" DiffE \<open>T \<subseteq> S - {a}\<close> mult_not_zero singletonI subsetCE \<open>a \<in> S\<close>)
finally show ?thesis
using \<open>0 \<notin> S\<close> \<open>a \<in> S\<close> by auto
@@ -13707,12 +13707,12 @@
if "x \<in> S" for x
proof -
have "a \<bullet> x = (\<Sum>y\<in>S. if y = x then y \<bullet> a else 0)"
- by (simp add: \<open>finite S\<close> inner_commute setsum.delta that)
+ by (simp add: \<open>finite S\<close> inner_commute sum.delta that)
also have "... = (\<Sum>b\<in>S. b \<bullet> a * (b \<bullet> x) / (b \<bullet> b))"
- apply (rule setsum.cong [OF refl], simp)
+ apply (rule sum.cong [OF refl], simp)
by (meson S orthogonal_def pairwise_def that)
finally show ?thesis
- by (simp add: orthogonal_def algebra_simps inner_setsum_left)
+ by (simp add: orthogonal_def algebra_simps inner_sum_left)
qed
then show ?thesis
using orthogonal_to_span orthogonal_commute x by blast
@@ -13746,7 +13746,7 @@
using spanU by simp
also have "... = span (insert a (S \<union> T))"
apply (rule eq_span_insert_eq)
- apply (simp add: a'_def span_neg span_setsum span_clauses(1) span_mul)
+ apply (simp add: a'_def span_neg span_sum span_clauses(1) span_mul)
done
also have "... = span (S \<union> insert a T)"
by simp
@@ -13866,7 +13866,7 @@
by (simp add: Gram_Schmidt_step \<open>pairwise orthogonal T\<close> \<open>span T = span S\<close> orthogonal_commute that)
show ?thesis
apply (rule_tac y = "?a" and z = "x - ?a" in that)
- apply (meson \<open>T \<subseteq> span S\<close> span_mul span_setsum subsetCE)
+ apply (meson \<open>T \<subseteq> span S\<close> span_mul span_sum subsetCE)
apply (fact orth, simp)
done
qed
--- a/src/HOL/Analysis/Derivative.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Derivative.thy Mon Oct 17 14:37:32 2016 +0200
@@ -2195,13 +2195,13 @@
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::banach"
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) {..<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 (sum (\<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)"
unfolding sums_def
apply (rule has_derivative_sequence[OF assms(1) _ assms(3)])
- apply (metis assms(2) has_derivative_setsum)
+ apply (metis assms(2) has_derivative_sum)
using assms(4-5)
unfolding sums_def
apply auto
@@ -2224,7 +2224,7 @@
{
fix n :: nat and x h :: 'a assume nx: "n \<ge> N" "x \<in> s"
have "norm ((\<Sum>i<n. f' i x * h) - g' x * h) = norm ((\<Sum>i<n. f' i x) - g' x) * norm h"
- by (simp add: norm_mult [symmetric] ring_distribs setsum_distrib_right)
+ by (simp add: norm_mult [symmetric] ring_distribs sum_distrib_right)
also from N[OF nx] have "norm ((\<Sum>i<n. f' i x) - g' x) \<le> e" by simp
hence "norm ((\<Sum>i<n. f' i x) - g' x) * norm h \<le> e * norm h"
by (intro mult_right_mono) simp_all
@@ -2591,7 +2591,7 @@
"\<forall>\<^sub>F n in sequentially. ((\<lambda>x. \<Sum>i<n. ?e i x) \<longlongrightarrow> A) ?F"
by eventually_elim
(auto intro!: tendsto_eq_intros
- simp: power_0_left if_distrib cond_application_beta setsum.delta
+ simp: power_0_left if_distrib cond_application_beta sum.delta
cong: if_cong)
ultimately
have [tendsto_intros]: "((\<lambda>x. \<Sum>i. ?e i x) \<longlongrightarrow> A) ?F"
--- a/src/HOL/Analysis/Determinants.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Determinants.thy Mon Oct 17 14:37:32 2016 +0200
@@ -13,7 +13,7 @@
subsection \<open>Trace\<close>
definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a"
- where "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
+ where "trace A = sum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
lemma trace_0: "trace (mat 0) = 0"
by (simp add: trace_def mat_def)
@@ -22,14 +22,14 @@
by (simp add: trace_def mat_def)
lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
- by (simp add: trace_def setsum.distrib)
+ by (simp add: trace_def sum.distrib)
lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
- by (simp add: trace_def setsum_subtractf)
+ by (simp add: trace_def sum_subtractf)
lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)"
apply (simp add: trace_def matrix_matrix_mult_def)
- apply (subst setsum.commute)
+ apply (subst sum.commute)
apply (simp add: mult.commute)
done
@@ -37,7 +37,7 @@
definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
"det A =
- setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
+ sum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set))
{p. p permutes (UNIV :: 'n set)}"
text \<open>A few general lemmas we need below.\<close>
@@ -94,8 +94,8 @@
}
then show ?thesis
unfolding det_def
- apply (subst setsum_permutations_inverse)
- apply (rule setsum.cong)
+ apply (subst sum_permutations_inverse)
+ apply (rule sum.cong)
apply (rule refl)
apply blast
done
@@ -128,7 +128,7 @@
}
then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
by blast
- from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
+ from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
unfolding det_def by (simp add: sign_id)
qed
@@ -159,7 +159,7 @@
}
then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
by blast
- from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
+ from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
unfolding det_def by (simp add: sign_id)
qed
@@ -189,7 +189,7 @@
}
then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
by blast
- from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
+ from sum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
unfolding det_def by (simp add: sign_id)
qed
@@ -226,9 +226,9 @@
fixes A :: "'a::comm_ring_1^'n^'n"
assumes p: "p permutes (UNIV :: 'n::finite set)"
shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
- apply (simp add: det_def setsum_distrib_left mult.assoc[symmetric])
+ apply (simp add: det_def sum_distrib_left mult.assoc[symmetric])
apply (subst sum_permutations_compose_right[OF p])
-proof (rule setsum.cong)
+proof (rule sum.cong)
let ?U = "UNIV :: 'n set"
let ?PU = "{p. p permutes ?U}"
fix q
@@ -303,7 +303,7 @@
shows "det A = 0"
using r
apply (simp add: row_def det_def vec_eq_iff)
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply (auto simp: sign_nz)
done
@@ -321,8 +321,8 @@
shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
- unfolding det_def vec_lambda_beta setsum.distrib[symmetric]
-proof (rule setsum.cong)
+ unfolding det_def vec_lambda_beta sum.distrib[symmetric]
+proof (rule sum.cong)
let ?U = "UNIV :: 'n set"
let ?pU = "{p. p permutes ?U}"
let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
@@ -372,8 +372,8 @@
fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
- unfolding det_def vec_lambda_beta setsum_distrib_left
-proof (rule setsum.cong)
+ unfolding det_def vec_lambda_beta sum_distrib_left
+proof (rule sum.cong)
let ?U = "UNIV :: 'n set"
let ?pU = "{p. p permutes ?U}"
let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
@@ -535,15 +535,15 @@
lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)"
by (rule iffD1[OF vec_lambda_unique]) vector
-lemma det_linear_row_setsum:
+lemma det_linear_row_sum:
assumes fS: "finite S"
- shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) =
- setsum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S"
+ shows "det ((\<chi> i. if i = k then sum (a i) S else c i)::'a::comm_ring_1^'n^'n) =
+ sum (\<lambda>j. det ((\<chi> i. if i = k then a i j else c i)::'a^'n^'n)) S"
proof (induct rule: finite_induct[OF fS])
case 1
then show ?case
apply simp
- unfolding setsum.empty det_row_0[of k]
+ unfolding sum.empty det_row_0[of k]
apply rule
done
next
@@ -577,11 +577,11 @@
qed
-lemma det_linear_rows_setsum_lemma:
+lemma det_linear_rows_sum_lemma:
assumes fS: "finite S"
and fT: "finite T"
- shows "det ((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) =
- setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
+ shows "det ((\<chi> i. if i \<in> T then sum (a i) S else c i):: 'a::comm_ring_1^'n^'n) =
+ sum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
{f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
using fT
proof (induct T arbitrary: a c set: finite)
@@ -604,48 +604,48 @@
by simp
from \<open>z \<notin> T\<close> have nz: "\<And>i. i \<in> T \<Longrightarrow> i = z \<longleftrightarrow> False"
by auto
- have "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
- det (\<chi> i. if i = z then setsum (a i) S else if i \<in> T then setsum (a i) S else c i)"
+ have "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) =
+ det (\<chi> i. if i = z then sum (a i) S else if i \<in> T then sum (a i) S else c i)"
unfolding insert_iff thif ..
- also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then setsum (a i) S else if i = z then a i j else c i))"
- unfolding det_linear_row_setsum[OF fS]
+ also have "\<dots> = (\<Sum>j\<in>S. det (\<chi> i. if i \<in> T then sum (a i) S else if i = z then a i j else c i))"
+ unfolding det_linear_row_sum[OF fS]
apply (subst thif2)
using nz
apply (simp cong del: if_weak_cong cong add: if_cong)
done
finally have tha:
- "det (\<chi> i. if i \<in> insert z T then setsum (a i) S else c i) =
+ "det (\<chi> i. if i \<in> insert z T then sum (a i) S else c i) =
(\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
else if i = z then a i j
else c i))"
- unfolding insert.hyps unfolding setsum.cartesian_product by blast
+ unfolding insert.hyps unfolding sum.cartesian_product by blast
show ?case unfolding tha
using \<open>z \<notin> T\<close>
- by (intro setsum.reindex_bij_witness[where i="?k" and j="?h"])
+ by (intro sum.reindex_bij_witness[where i="?k" and j="?h"])
(auto intro!: cong[OF refl[of det]] simp: vec_eq_iff)
qed
-lemma det_linear_rows_setsum:
+lemma det_linear_rows_sum:
fixes S :: "'n::finite set"
assumes fS: "finite S"
- shows "det (\<chi> i. setsum (a i) S) =
- setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}"
+ shows "det (\<chi> i. sum (a i) S) =
+ sum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}"
proof -
have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)"
by vector
- from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite]
+ from det_linear_rows_sum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite]
show ?thesis by simp
qed
-lemma matrix_mul_setsum_alt:
+lemma matrix_mul_sum_alt:
fixes A B :: "'a::comm_ring_1^'n^'n"
- shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
- by (vector matrix_matrix_mult_def setsum_component)
+ shows "A ** B = (\<chi> i. sum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
+ by (vector matrix_matrix_mult_def sum_component)
lemma det_rows_mul:
"det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
-proof (simp add: det_def setsum_distrib_left cong add: setprod.cong, rule setsum.cong)
+proof (simp add: det_def sum_distrib_left cong add: setprod.cong, rule sum.cong)
let ?U = "UNIV :: 'n set"
let ?PU = "{p. p permutes ?U}"
fix p
@@ -735,11 +735,11 @@
by blast
let ?s = "\<lambda>p. of_int (sign p)"
let ?f = "\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))"
- have "(setsum (\<lambda>q. ?s q *
+ have "(sum (\<lambda>q. ?s q *
(\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) =
- (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)"
+ (sum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)"
unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f]
- proof (rule setsum.cong)
+ proof (rule sum.cong)
fix q
assume qU: "q \<in> ?PU"
then have q: "q permutes ?U"
@@ -769,14 +769,14 @@
by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose)
qed rule
}
- then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
- unfolding det_def setsum_product
- by (rule setsum.cong [OF refl])
- have "det (A**B) = setsum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
- unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU]
+ then have th2: "sum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
+ unfolding det_def sum_product
+ by (rule sum.cong [OF refl])
+ have "det (A**B) = sum (\<lambda>f. det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
+ unfolding matrix_mul_sum_alt det_linear_rows_sum[OF fU]
by simp
- also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
- using setsum.mono_neutral_cong_left[OF fF PUF zth, symmetric]
+ also have "\<dots> = sum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
+ using sum.mono_neutral_cong_left[OF fF PUF zth, symmetric]
unfolding det_rows_mul by auto
finally show ?thesis unfolding th2 .
qed
@@ -812,7 +812,7 @@
let ?U = "UNIV :: 'n set"
have fU: "finite ?U"
by simp
- from H obtain c i where c: "setsum (\<lambda>i. c i *s row i A) ?U = 0"
+ from H obtain c i where c: "sum (\<lambda>i. c i *s row i A) ?U = 0"
and iU: "i \<in> ?U"
and ci: "c i \<noteq> 0"
unfolding invertible_righ_inverse
@@ -824,8 +824,8 @@
apply simp
done
from c ci
- have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
- unfolding setsum.remove[OF fU iU] setsum_cmul
+ have thr0: "- row i A = sum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
+ unfolding sum.remove[OF fU iU] sum_cmul
apply -
apply (rule vector_mul_lcancel_imp[OF ci])
apply (auto simp add: field_simps)
@@ -834,7 +834,7 @@
done
have thr: "- row i A \<in> span {row j A| j. j \<noteq> i}"
unfolding thr0
- apply (rule span_setsum)
+ apply (rule span_sum)
apply simp
apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+
apply (rule span_superset)
@@ -855,7 +855,7 @@
lemma cramer_lemma_transpose:
fixes A:: "real^'n^'n"
and x :: "real^'n"
- shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
+ shows "det ((\<chi> i. if i = k then sum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
else row i A)::real^'n^'n) = x$k * det A"
(is "?lhs = ?rhs")
proof -
@@ -876,14 +876,14 @@
by simp
have thd0: "det (\<chi> i. if i = k then row k A + (\<Sum>i \<in> ?Uk. x $ i *s row i A) else row i A) = det A"
apply (rule det_row_span)
- apply (rule span_setsum)
+ apply (rule span_sum)
apply (rule span_mul [where 'a="real^'n", folded scalar_mult_eq_scaleR])+
apply (rule span_superset)
apply auto
done
show "?lhs = x$k * det A"
apply (subst U)
- unfolding setsum.insert[OF fUk kUk]
+ unfolding sum.insert[OF fUk kUk]
apply (subst th00)
unfolding add.assoc
apply (subst det_row_add)
@@ -900,8 +900,8 @@
shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: real^'n^'n) = x$k * det A"
proof -
let ?U = "UNIV :: 'n set"
- have *: "\<And>c. setsum (\<lambda>i. c i *s row i (transpose A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
- by (auto simp add: row_transpose intro: setsum.cong)
+ have *: "\<And>c. sum (\<lambda>i. c i *s row i (transpose A)) ?U = sum (\<lambda>i. c i *s column i A) ?U"
+ by (auto simp add: row_transpose intro: sum.cong)
show ?thesis
unfolding matrix_mult_vsum
unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric]
@@ -995,7 +995,7 @@
simplified matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
have "?A$i$j = ?m1 $ i $ j"
by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def
- th0 setsum.delta[OF fU] mat_def axis_def)
+ th0 sum.delta[OF fU] mat_def axis_def)
}
then have "orthogonal_matrix ?mf"
unfolding orthogonal_matrix
@@ -1219,7 +1219,7 @@
have f12: "finite {2::2}" "1 \<notin> {2::2}" by auto
show ?thesis
unfolding det_def UNIV_2
- unfolding setsum_over_permutations_insert[OF f12]
+ unfolding sum_over_permutations_insert[OF f12]
unfolding permutes_sing
by (simp add: sign_swap_id sign_id swap_id_eq)
qed
@@ -1240,8 +1240,8 @@
show ?thesis
unfolding det_def UNIV_3
- unfolding setsum_over_permutations_insert[OF f123]
- unfolding setsum_over_permutations_insert[OF f23]
+ unfolding sum_over_permutations_insert[OF f123]
+ unfolding sum_over_permutations_insert[OF f23]
unfolding permutes_sing
by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq)
qed
--- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Oct 17 14:37:32 2016 +0200
@@ -182,15 +182,15 @@
have "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
(\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p - s. content k *\<^sub>R f x)"
using p(1)[THEN tagged_division_ofD(1)]
- by (safe intro!: setsum.union_inter_neutral) (auto simp: s_def T_def)
+ by (safe intro!: sum.union_inter_neutral) (auto simp: s_def T_def)
also have "(\<Sum>(x, k)\<in>?T ` (p \<inter> s). content k *\<^sub>R f x) = (\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k))"
- proof (subst setsum.reindex_nontrivial, safe)
+ proof (subst sum.reindex_nontrivial, safe)
fix x1 x2 k assume 1: "(x1, k) \<in> p" "(x1, k) \<in> s" and 2: "(x2, k) \<in> p" "(x2, k) \<in> s"
and eq: "content k *\<^sub>R f (T X k) \<noteq> 0"
with tagged_division_ofD(5)[OF p(1), of x1 k x2 k] tagged_division_ofD(4)[OF p(1), of x1 k]
show "x1 = x2"
by (auto simp: content_eq_0_interior)
- qed (use p in \<open>auto intro!: setsum.cong\<close>)
+ qed (use p in \<open>auto intro!: sum.cong\<close>)
finally have eq: "(\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) =
(\<Sum>(x, k)\<in>p \<inter> s. content k *\<^sub>R f (T X k)) + (\<Sum>(x, k)\<in>p - s. content k *\<^sub>R f x)" .
@@ -270,11 +270,11 @@
using \<open>0 < e\<close> by (simp add: split_beta)
qed (use \<open>a < b\<close> in auto)
also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * (b - a))"
- by (simp add: setsum_distrib_right split_beta')
+ by (simp add: sum_distrib_right split_beta')
also have "\<dots> \<le> (\<Sum>(x, k)\<in>p \<inter> s. content k * (f (T ?F k) - f (T ?E k)))"
- using parts(3) by (auto intro!: setsum_mono mult_left_mono diff_mono)
+ using parts(3) by (auto intro!: sum_mono mult_left_mono diff_mono)
also have "\<dots> = (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?F k)) - (\<Sum>(x, k)\<in>p \<inter> s. content k * f (T ?E k))"
- by (auto intro!: setsum.cong simp: field_simps setsum_subtractf[symmetric])
+ by (auto intro!: sum.cong simp: field_simps sum_subtractf[symmetric])
also have "\<dots> = (\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x)"
by (subst (1 2) parts) auto
also have "\<dots> \<le> norm ((\<Sum>(x, k)\<in>?B. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>?A. content k *\<^sub>R f x))"
@@ -401,13 +401,13 @@
proof (rule has_integral_monotone_convergence_increasing)
let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
- using union.IH by (auto intro!: has_integral_setsum simp del: Int_iff)
+ using union.IH by (auto intro!: has_integral_sum simp del: Int_iff)
show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
- by (intro setsum_mono2) auto
+ by (intro sum_mono2) auto
from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
by (auto simp add: disjoint_family_on_def)
show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
- apply (auto simp: * setsum.If_cases Iio_Int_singleton)
+ apply (auto simp: * sum.If_cases Iio_Int_singleton)
apply (rule_tac k="Suc xa" in LIMSEQ_offset)
apply simp
done
@@ -757,7 +757,7 @@
shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
proof -
have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
- using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
+ using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
by (simp add: fun_eq_iff euclidean_representation)
also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
@@ -804,7 +804,7 @@
shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV"
proof -
have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
- using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto
+ using f by (intro has_integral_sum finite_Basis ballI has_integral_scaleR_left has_integral_integral_lebesgue_real) auto
also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
by (simp add: fun_eq_iff euclidean_representation)
also have "(\<Sum>b\<in>Basis. integral\<^sup>L lebesgue (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lebesgue f"
@@ -1245,10 +1245,10 @@
using \<open>0 < B\<close> \<open>\<D>' \<subseteq> \<D>\<close> that vu_pos by fastforce
have "{} \<notin> \<D>'"
using cbox \<open>\<D>' \<subseteq> \<D>\<close> interior_empty by blast
- have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> setsum (measure lebesgue o fbx) \<D>'"
- by (rule setsum_image_le [OF \<open>finite \<D>'\<close>]) (force simp: fbx_def)
+ have "(\<Sum>k\<in>fbx`\<D>'. measure lebesgue k) \<le> sum (measure lebesgue o fbx) \<D>'"
+ by (rule sum_image_le [OF \<open>finite \<D>'\<close>]) (force simp: fbx_def)
also have "\<dots> \<le> (\<Sum>X\<in>\<D>'. (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X)"
- proof (rule setsum_mono)
+ proof (rule sum_mono)
fix X assume "X \<in> \<D>'"
then have "X \<in> \<D>" using \<open>\<D>' \<subseteq> \<D>\<close> by blast
then have ufvf: "cbox (uf X) (vf X) = X"
@@ -1281,8 +1281,8 @@
by (subst (3) ufvf[symmetric]) simp
finally show "(measure lebesgue \<circ> fbx) X \<le> (2 * B * DIM('M)) ^ DIM('N) * measure lebesgue X" .
qed
- also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * setsum (measure lebesgue) \<D>'"
- by (simp add: setsum_distrib_left)
+ also have "\<dots> = (2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>'"
+ by (simp add: sum_distrib_left)
also have "\<dots> \<le> e / 2"
proof -
have div: "\<D>' division_of \<Union>\<D>'"
@@ -1306,9 +1306,9 @@
qed
finally show "\<Union>\<D>' \<subseteq> T" .
qed
- have "setsum (measure lebesgue) \<D>' = setsum content \<D>'"
- using \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: setsum.cong)
- then have "(2 * B * DIM('M)) ^ DIM('N) * setsum (measure lebesgue) \<D>' =
+ have "sum (measure lebesgue) \<D>' = sum content \<D>'"
+ using \<open>\<D>' \<subseteq> \<D>\<close> cbox by (force intro: sum.cong)
+ then have "(2 * B * DIM('M)) ^ DIM('N) * sum (measure lebesgue) \<D>' =
(2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue (\<Union>\<D>')"
using content_division [OF div] by auto
also have "\<dots> \<le> (2 * B * real DIM('M)) ^ DIM('N) * measure lebesgue T"
@@ -1322,7 +1322,7 @@
qed
finally show ?thesis .
qed
- then have e2: "setsum (measure lebesgue) \<G> \<le> e / 2" if "\<G> \<subseteq> fbx ` \<D>" "finite \<G>" for \<G>
+ then have e2: "sum (measure lebesgue) \<G> \<le> e / 2" if "\<G> \<subseteq> fbx ` \<D>" "finite \<G>" for \<G>
by (metis finite_subset_image that)
show "\<exists>W\<in>lmeasurable. f ` S \<subseteq> W \<and> measure lebesgue W < e"
proof (intro bexI conjI)
@@ -1339,7 +1339,7 @@
have "norm (y - zf X) \<le> (\<Sum>i\<in>Basis. \<bar>(y - zf X) \<bullet> i\<bar>)"
by (rule norm_le_l1)
also have "\<dots> \<le> real DIM('M) * prj1 (vf X - uf X)"
- proof (rule setsum_bounded_above)
+ proof (rule sum_bounded_above)
fix j::'M assume j: "j \<in> Basis"
show "\<bar>(y - zf X) \<bullet> j\<bar> \<le> prj1 (vf X - uf X)"
using yin zin j
@@ -1533,7 +1533,7 @@
lemma absolutely_integrable_bounded_variation:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes f: "f absolutely_integrable_on UNIV"
- obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
+ obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
proof (rule that[of "integral UNIV (\<lambda>x. norm (f x))"]; safe)
fix d :: "'a set set" assume d: "d division_of \<Union>d"
have *: "k \<in> d \<Longrightarrow> f absolutely_integrable_on k" for k
@@ -1541,11 +1541,11 @@
note d' = division_ofD[OF d]
have "(\<Sum>k\<in>d. norm (integral k f)) = (\<Sum>k\<in>d. norm (LINT x:k|lebesgue. f x))"
- by (intro setsum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *)
+ by (intro sum.cong refl arg_cong[where f=norm] set_lebesgue_integral_eq_integral(2)[symmetric] *)
also have "\<dots> \<le> (\<Sum>k\<in>d. LINT x:k|lebesgue. norm (f x))"
- by (intro setsum_mono set_integral_norm_bound *)
+ by (intro sum_mono set_integral_norm_bound *)
also have "\<dots> = (\<Sum>k\<in>d. integral k (\<lambda>x. norm (f x)))"
- by (intro setsum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *)
+ by (intro sum.cong refl set_lebesgue_integral_eq_integral(2) set_integrable_norm *)
also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))"
using integrable_on_subdivision[OF d] assms f unfolding absolutely_integrable_on_def
by (subst integral_combine_division_topdown[OF _ d]) auto
@@ -1556,20 +1556,20 @@
qed
lemma helplemma:
- assumes "setsum (\<lambda>x. norm (f x - g x)) s < e"
+ assumes "sum (\<lambda>x. norm (f x - g x)) s < e"
and "finite s"
- shows "\<bar>setsum (\<lambda>x. norm(f x)) s - setsum (\<lambda>x. norm(g x)) s\<bar> < e"
- unfolding setsum_subtractf[symmetric]
- apply (rule le_less_trans[OF setsum_abs])
+ shows "\<bar>sum (\<lambda>x. norm(f x)) s - sum (\<lambda>x. norm(g x)) s\<bar> < e"
+ unfolding sum_subtractf[symmetric]
+ apply (rule le_less_trans[OF sum_abs])
apply (rule le_less_trans[OF _ assms(1)])
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply (rule norm_triangle_ineq3)
done
lemma bounded_variation_absolutely_integrable_interval:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes f: "f integrable_on cbox a b"
- and *: "\<forall>d. d division_of (cbox a b) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
+ and *: "\<forall>d. d division_of (cbox a b) \<longrightarrow> sum (\<lambda>k. norm(integral k f)) d \<le> B"
shows "f absolutely_integrable_on cbox a b"
proof -
let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}"
@@ -1749,7 +1749,7 @@
done
qed
have sum_p': "(\<Sum>(x, k)\<in>p'. norm (integral k f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
- apply (subst setsum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
+ apply (subst sum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
unfolding norm_eq_zero
apply (rule integral_null)
apply (simp add: content_eq_0_interior)
@@ -1775,7 +1775,7 @@
(\<lambda>(k,l). k \<inter> l) ` {(k,l)|k l. k \<in> d \<and> l \<in> snd ` p}"
by auto
have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))"
- proof (rule setsum_mono, goal_cases)
+ proof (rule sum_mono, goal_cases)
case k: (1 k)
from d'(4)[OF this] guess u v by (elim exE) note uv=this
define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and> cbox u v \<inter> l \<noteq> {}}"
@@ -1786,16 +1786,16 @@
apply (rule division_of_tagged_division[OF p(1)])
apply (rule uvab)
done
- then have "norm (integral k f) \<le> setsum (\<lambda>k. norm (integral k f)) d'"
+ then have "norm (integral k f) \<le> sum (\<lambda>k. norm (integral k f)) d'"
unfolding uv
apply (subst integral_combine_division_topdown[of _ _ d'])
apply (rule integrable_on_subcbox[OF assms(1) uvab])
apply assumption
- apply (rule setsum_norm_le)
+ apply (rule sum_norm_le)
apply auto
done
also have "\<dots> = (\<Sum>k\<in>{k \<inter> l |l. l \<in> snd ` p}. norm (integral k f))"
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
apply (subst Setcompr_eq_image)
apply (rule finite_imageI)+
apply fact
@@ -1813,7 +1813,7 @@
qed
also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (k \<inter> l) f))"
unfolding Setcompr_eq_image
- apply (rule setsum.reindex_nontrivial [unfolded o_def])
+ apply (rule sum.reindex_nontrivial [unfolded o_def])
apply (rule finite_imageI)
apply (rule p')
proof goal_cases
@@ -1845,7 +1845,7 @@
unfolding split_def ..
also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
unfolding *
- apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def])
+ apply (rule sum.reindex_nontrivial [symmetric, unfolded o_def])
apply (rule finite_product_dependent)
apply fact
apply (rule finite_imageI)
@@ -1886,7 +1886,7 @@
qed
also have "\<dots> = (\<Sum>(x, k)\<in>p'. norm (integral k f))"
unfolding sum_p'
- apply (rule setsum.mono_neutral_right)
+ apply (rule sum.mono_neutral_right)
apply (subst *)
apply (rule finite_imageI[OF finite_product_dependent])
apply fact
@@ -1929,7 +1929,7 @@
note pdfin = finite_cartesian_product[OF p'(1) d'(1)]
have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x, k)\<in>?S. \<bar>content k\<bar> * norm (f x))"
unfolding norm_scaleR
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
apply (subst *)
apply (rule finite_imageI)
apply fact
@@ -1943,7 +1943,7 @@
done
also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
unfolding *
- apply (subst setsum.reindex_nontrivial)
+ apply (subst sum.reindex_nontrivial)
apply fact
unfolding split_paired_all
unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject
@@ -1981,7 +1981,7 @@
apply (rule p')
apply rule
apply (rule d')
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
unfolding split_paired_all split_conv
proof -
@@ -1990,7 +1990,7 @@
note xl = p'(2-4)[OF this]
from this(3) guess u v by (elim exE) note uv=this
have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
apply (drule d'(4))
apply safe
@@ -1999,9 +1999,9 @@
apply (subst abs_of_nonneg)
apply auto
done
- also have "\<dots> = setsum content {k \<inter> cbox u v| k. k \<in> d}"
+ also have "\<dots> = sum content {k \<inter> cbox u v| k. k \<in> d}"
unfolding Setcompr_eq_image
- apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric])
+ apply (rule sum.reindex_nontrivial [unfolded o_def, symmetric])
apply (rule d')
proof goal_cases
case prems: (1 k y)
@@ -2019,8 +2019,8 @@
finally show ?case
unfolding uv Int_interval content_eq_0_interior ..
qed
- also have "\<dots> = setsum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
- apply (rule setsum.mono_neutral_right)
+ also have "\<dots> = sum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
+ apply (rule sum.mono_neutral_right)
unfolding Setcompr_eq_image
apply (rule finite_imageI)
apply (rule d')
@@ -2040,7 +2040,7 @@
by auto
qed
finally show "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar> * norm (f x)) = content l *\<^sub>R norm (f x)"
- unfolding setsum_distrib_right[symmetric] real_scaleR_def
+ unfolding sum_distrib_right[symmetric] real_scaleR_def
apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
using xl(2)[unfolded uv]
unfolding uv
@@ -2056,7 +2056,7 @@
lemma bounded_variation_absolutely_integrable:
fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "f integrable_on UNIV"
- and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
+ and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> sum (\<lambda>k. norm (integral k f)) d \<le> B"
shows "f absolutely_integrable_on UNIV"
proof (rule absolutely_integrable_onI, fact, rule)
let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (\<Union>d)}"
@@ -2083,7 +2083,7 @@
using f_int[of a b] unfolding absolutely_integrable_on_def by auto
next
case prems: (2 e)
- have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
+ have "\<exists>y\<in>sum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
proof (rule ccontr)
assume "\<not> ?thesis"
then have "?S \<le> ?S - e"
@@ -2092,7 +2092,7 @@
using prems by auto
qed
then obtain K where *: "\<exists>x\<in>{d. d division_of \<Union>d}. K = (\<Sum>k\<in>x. norm (integral k f))"
- "SUPREMUM {d. d division_of \<Union>d} (setsum (\<lambda>k. norm (integral k f))) - e < K"
+ "SUPREMUM {d. d division_of \<Union>d} (sum (\<lambda>k. norm (integral k f))) - e < K"
by (auto simp add: image_iff not_le)
from this(1) obtain d where "d division_of \<Union>d"
and "K = (\<Sum>k\<in>d. norm (integral k f))"
@@ -2118,8 +2118,8 @@
apply (rule d(2))
proof goal_cases
case 1
- have "(\<Sum>k\<in>d. norm (integral k f)) \<le> setsum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
- apply (intro setsum_mono)
+ have "(\<Sum>k\<in>d. norm (integral k f)) \<le> sum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
+ apply (intro sum_mono)
subgoal for k
using d'(4)[of k] f_int
by (auto simp: absolutely_integrable_on_def integral_norm_bound_integral)
@@ -2188,7 +2188,7 @@
using d1(2)[rule_format,OF conjI[OF p(1,2)]]
by (simp only: real_norm_def)
show "(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) = (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x))"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
unfolding split_paired_all split_conv
apply (drule tagged_division_ofD(4)[OF p(1)])
@@ -2198,7 +2198,7 @@
done
show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
using partial_division_of_tagged_division[of p "cbox a b"] p(1)
- apply (subst setsum.over_tagged_division_lemma[OF p(1)])
+ apply (subst sum.over_tagged_division_lemma[OF p(1)])
apply (simp add: content_eq_0_interior)
apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
apply (auto simp: tagged_partial_division_of_def)
@@ -2230,10 +2230,10 @@
using integrable_bounded_linear[of h lebesgue "\<lambda>x. indicator s x *\<^sub>R f x"]
by (simp add: linear_simps[of h])
-lemma absolutely_integrable_setsum:
+lemma absolutely_integrable_sum:
fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
assumes "finite t" and "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s"
- shows "(\<lambda>x. setsum (\<lambda>a. f a x) t) absolutely_integrable_on s"
+ shows "(\<lambda>x. sum (\<lambda>a. f a x) t) absolutely_integrable_on s"
using assms(1,2) by induct auto
lemma absolutely_integrable_integrable_bound:
--- a/src/HOL/Analysis/Euclidean_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Euclidean_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -53,9 +53,9 @@
lemma (in euclidean_space) SOME_Basis: "(SOME i. i \<in> Basis) \<in> Basis"
by (metis ex_in_conv nonempty_Basis someI_ex)
-lemma (in euclidean_space) inner_setsum_left_Basis[simp]:
+lemma (in euclidean_space) inner_sum_left_Basis[simp]:
"b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b"
- by (simp add: inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.If_cases)
+ by (simp add: inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.If_cases)
lemma (in euclidean_space) euclidean_eqI:
assumes b: "\<And>b. b \<in> Basis \<Longrightarrow> inner x b = inner y b" shows "x = y"
@@ -70,16 +70,16 @@
"x = y \<longleftrightarrow> (\<forall>b\<in>Basis. inner x b = inner y b)"
by (auto intro: euclidean_eqI)
-lemma (in euclidean_space) euclidean_representation_setsum:
+lemma (in euclidean_space) euclidean_representation_sum:
"(\<Sum>i\<in>Basis. f i *\<^sub>R i) = b \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)"
by (subst euclidean_eq_iff) simp
-lemma (in euclidean_space) euclidean_representation_setsum':
+lemma (in euclidean_space) euclidean_representation_sum':
"b = (\<Sum>i\<in>Basis. f i *\<^sub>R i) \<longleftrightarrow> (\<forall>i\<in>Basis. f i = inner b i)"
- by (auto simp add: euclidean_representation_setsum[symmetric])
+ by (auto simp add: euclidean_representation_sum[symmetric])
lemma (in euclidean_space) euclidean_representation: "(\<Sum>b\<in>Basis. inner x b *\<^sub>R b) = x"
- unfolding euclidean_representation_setsum by simp
+ unfolding euclidean_representation_sum by simp
lemma (in euclidean_space) choice_Basis_iff:
fixes P :: "'a \<Rightarrow> real \<Rightarrow> bool"
@@ -96,15 +96,15 @@
shows "(\<forall>i\<in>Basis. \<exists>x\<in>A. P i x) \<longleftrightarrow> (\<exists>x. \<forall>i\<in>Basis. inner x i \<in> A \<and> P i (inner x i))"
by (simp add: choice_Basis_iff Bex_def)
-lemma (in euclidean_space) euclidean_representation_setsum_fun:
+lemma (in euclidean_space) euclidean_representation_sum_fun:
"(\<lambda>x. \<Sum>b\<in>Basis. inner (f x) b *\<^sub>R b) = f"
- by (rule ext) (simp add: euclidean_representation_setsum)
+ by (rule ext) (simp add: euclidean_representation_sum)
lemma euclidean_isCont:
assumes "\<And>b. b \<in> Basis \<Longrightarrow> isCont (\<lambda>x. (inner (f x) b) *\<^sub>R b) x"
shows "isCont f x"
- apply (subst euclidean_representation_setsum_fun [symmetric])
- apply (rule isCont_setsum)
+ apply (subst euclidean_representation_sum_fun [symmetric])
+ apply (rule isCont_sum)
apply (blast intro: assms)
done
@@ -115,16 +115,16 @@
by (meson DIM_positive Suc_leI)
-lemma setsum_inner_Basis_scaleR [simp]:
+lemma sum_inner_Basis_scaleR [simp]:
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_vector"
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) *\<^sub>R f i) = f b"
- by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms]
- assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral)
+ by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms]
+ assms inner_not_same_Basis comm_monoid_add_class.sum.neutral)
-lemma setsum_inner_Basis_eq [simp]:
+lemma sum_inner_Basis_eq [simp]:
assumes "b \<in> Basis" shows "(\<Sum>i\<in>Basis. (inner i b) * f i) = f b"
- by (simp add: comm_monoid_add_class.setsum.remove [OF finite_Basis assms]
- assms inner_not_same_Basis comm_monoid_add_class.setsum.neutral)
+ by (simp add: comm_monoid_add_class.sum.remove [OF finite_Basis assms]
+ assms inner_not_same_Basis comm_monoid_add_class.sum.neutral)
subsection \<open>Subclass relationships\<close>
@@ -188,15 +188,15 @@
definition
"Basis = (\<lambda>u. (u, 0)) ` Basis \<union> (\<lambda>v. (0, v)) ` Basis"
-lemma setsum_Basis_prod_eq:
+lemma sum_Basis_prod_eq:
fixes f::"('a*'b)\<Rightarrow>('a*'b)"
- shows "setsum f Basis = setsum (\<lambda>i. f (i, 0)) Basis + setsum (\<lambda>i. f (0, i)) Basis"
+ shows "sum f Basis = sum (\<lambda>i. f (i, 0)) Basis + sum (\<lambda>i. f (0, i)) Basis"
proof -
have "inj_on (\<lambda>u. (u::'a, 0::'b)) Basis" "inj_on (\<lambda>u. (0::'a, u::'b)) Basis"
by (auto intro!: inj_onI Pair_inject)
thus ?thesis
unfolding Basis_prod_def
- by (subst setsum.union_disjoint) (auto simp: Basis_prod_def setsum.reindex)
+ by (subst sum.union_disjoint) (auto simp: Basis_prod_def sum.reindex)
qed
instance proof
--- a/src/HOL/Analysis/Fashoda_Theorem.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Fashoda_Theorem.thy Mon Oct 17 14:37:32 2016 +0200
@@ -17,13 +17,13 @@
lemma interval_bij_affine:
"interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) +
(\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))"
- by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
- field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong)
+ by (auto simp: sum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
+ field_simps inner_simps add_divide_distrib[symmetric] intro!: sum.cong)
lemma continuous_interval_bij:
fixes a b :: "'a::euclidean_space"
shows "continuous (at x) (interval_bij (a, b) (u, v))"
- by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros)
+ by (auto simp add: divide_inverse interval_bij_def intro!: continuous_sum continuous_intros)
lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))"
apply(rule continuous_at_imp_continuous_on)
@@ -35,7 +35,7 @@
assumes "x \<in> cbox a b"
and "cbox u v \<noteq> {}"
shows "interval_bij (a, b) (u, v) x \<in> cbox u v"
- apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong)
+ apply (simp only: interval_bij_def split_conv mem_box inner_sum_left_Basis cong: ball_cong)
apply safe
proof -
fix i :: 'a
--- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Mon Oct 17 14:37:32 2016 +0200
@@ -373,10 +373,10 @@
assume "M \<le> m" "M \<le> n"
have "dist (X m) (X n) = setL2 (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
unfolding dist_vec_def ..
- also have "\<dots> \<le> setsum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
- by (rule setL2_le_setsum [OF zero_le_dist])
- also have "\<dots> < setsum (\<lambda>i::'n. ?s) UNIV"
- by (rule setsum_strict_mono, simp_all add: M \<open>M \<le> m\<close> \<open>M \<le> n\<close>)
+ also have "\<dots> \<le> sum (\<lambda>i. dist (X m $ i) (X n $ i)) UNIV"
+ by (rule setL2_le_sum [OF zero_le_dist])
+ also have "\<dots> < sum (\<lambda>i::'n. ?s) UNIV"
+ by (rule sum_strict_mono, simp_all add: M \<open>M \<le> m\<close> \<open>M \<le> n\<close>)
also have "\<dots> = r"
by simp
finally have "dist (X m) (X n) < r" .
@@ -449,7 +449,7 @@
instantiation vec :: (real_inner, finite) real_inner
begin
-definition "inner x y = setsum (\<lambda>i. inner (x$i) (y$i)) UNIV"
+definition "inner x y = sum (\<lambda>i. inner (x$i) (y$i)) UNIV"
instance proof
fix r :: real and x y z :: "'a ^ 'b"
@@ -458,16 +458,16 @@
by (simp add: inner_commute)
show "inner (x + y) z = inner x z + inner y z"
unfolding inner_vec_def
- by (simp add: inner_add_left setsum.distrib)
+ by (simp add: inner_add_left sum.distrib)
show "inner (scaleR r x) y = r * inner x y"
unfolding inner_vec_def
- by (simp add: setsum_distrib_left)
+ by (simp add: sum_distrib_left)
show "0 \<le> inner x x"
unfolding inner_vec_def
- by (simp add: setsum_nonneg)
+ by (simp add: sum_nonneg)
show "inner x x = 0 \<longleftrightarrow> x = 0"
unfolding inner_vec_def
- by (simp add: vec_eq_iff setsum_nonneg_eq_0_iff)
+ by (simp add: vec_eq_iff sum_nonneg_eq_0_iff)
show "norm x = sqrt (inner x x)"
unfolding inner_vec_def norm_vec_def setL2_def
by (simp add: power2_norm_eq_inner)
@@ -493,22 +493,22 @@
unfolding inner_vec_def
apply (cases "i = j")
apply clarsimp
- apply (subst setsum.remove [of _ j], simp_all)
- apply (rule setsum.neutral, simp add: axis_def)
- apply (rule setsum.neutral, simp add: axis_def)
+ apply (subst sum.remove [of _ j], simp_all)
+ apply (rule sum.neutral, simp add: axis_def)
+ apply (rule sum.neutral, simp add: axis_def)
done
-lemma setsum_single:
+lemma sum_single:
assumes "finite A" and "k \<in> A" and "f k = y"
assumes "\<And>i. i \<in> A \<Longrightarrow> i \<noteq> k \<Longrightarrow> f i = 0"
shows "(\<Sum>i\<in>A. f i) = y"
- apply (subst setsum.remove [OF assms(1,2)])
- apply (simp add: setsum.neutral assms(3,4))
+ apply (subst sum.remove [OF assms(1,2)])
+ apply (simp add: sum.neutral assms(3,4))
done
lemma inner_axis: "inner x (axis i y) = inner (x $ i) y"
unfolding inner_vec_def
- apply (rule_tac k=i in setsum_single)
+ apply (rule_tac k=i in sum_single)
apply simp_all
apply (simp add: axis_def)
done
--- a/src/HOL/Analysis/Gamma_Function.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Gamma_Function.thy Mon Oct 17 14:37:32 2016 +0200
@@ -407,7 +407,7 @@
moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all
moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n
using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def
- by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps)
+ by (subst (asm) abs_of_nonneg) (auto intro: sum_nonneg simp: divide_simps)
moreover note calculation
} note N = this
@@ -432,25 +432,25 @@
by (simp add: ln_Gamma_series_def algebra_simps)
also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) =
(\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
- by (simp add: setsum_diff)
+ by (simp add: sum_diff)
also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce
also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) =
(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn
- by (subst setsum_telescope'' [symmetric]) simp_all
+ by (subst sum_telescope'' [symmetric]) simp_all
also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N
- by (intro setsum_cong_Suc)
+ by (intro sum_cong_Suc)
(simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat)
also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k
using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps)
hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) =
(\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N
- by (intro setsum.cong) simp_all
- also note setsum.distrib [symmetric]
+ by (intro sum.cong) simp_all
+ also note sum.distrib [symmetric]
also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le>
(\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t
- by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
+ by (intro order.trans[OF norm_sum sum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
- by (simp add: setsum_distrib_left)
+ by (simp add: sum_distrib_left)
also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')"
@@ -508,7 +508,7 @@
from assms have "z \<noteq> 0" by (intro notI) auto
with assms have "exp (ln_Gamma_series z n) =
(of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"
- unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum)
+ unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_sum)
also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)"
by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp)
also have "... = (\<Prod>k=1..n. z + k) / fact n"
@@ -540,12 +540,12 @@
proof eventually_elim
fix n :: nat assume n: "n > 0"
have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"
- by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric],
+ by (subst atLeast0LessThan [symmetric], subst sum_shift_bounds_Suc_ivl [symmetric],
subst atLeastLessThanSuc_atLeastAtMost) simp_all
also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
- by (simp add: harm_def setsum_subtractf setsum_distrib_left divide_inverse)
+ by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse)
also from n have "\<dots> - ?g n = 0"
- by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat)
+ by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps Ln_of_nat)
finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
qed
show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
@@ -654,13 +654,13 @@
fix m :: nat
have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto
also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)"
- by (rule setsum.union_disjoint) auto
+ by (rule sum.union_disjoint) auto
also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
- by (intro setsum.reindex_cong[of "\<lambda>n. n + k"])
+ by (intro sum.reindex_cong[of "\<lambda>n. n + k"])
(simp, subst image_add_atLeastLessThan, auto)
finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan)
qed
- finally have "(\<lambda>a. setsum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. setsum f {..<m + k})"
+ finally have "(\<lambda>a. sum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. sum f {..<m + k})"
by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset)
thus ?thesis by (auto simp: summable_iff_convergent convergent_def)
qed
@@ -733,10 +733,10 @@
sums Digamma z" by (simp add: add_ac)
hence "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1))) -
(\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"
- by (simp add: sums_def setsum_subtractf)
+ by (simp add: sums_def sum_subtractf)
also have "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)))) =
(\<lambda>m. of_real (ln (m + 1)) :: 'a)"
- by (subst setsum_lessThan_telescope) simp_all
+ by (subst sum_lessThan_telescope) simp_all
finally show ?thesis by (rule Lim_transform) (insert lim, simp)
qed
@@ -944,7 +944,7 @@
using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
thus "uniformly_convergent_on (ball z d)
(\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
- by (subst (asm) setsum_distrib_left) simp
+ by (subst (asm) sum_distrib_left) simp
qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
(- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
@@ -2573,7 +2573,7 @@
proof -
have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) =
exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
- by (subst exp_setsum [symmetric]) (simp_all add: setsum_distrib_left)
+ by (subst exp_sum [symmetric]) (simp_all add: sum_distrib_left)
also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg)
also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
@@ -2654,9 +2654,9 @@
have "(\<lambda>n. \<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \<longlonglongrightarrow> ln_Gamma z + euler_mascheroni * z + ln z"
using ln_Gamma_series'_aux[OF False]
by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def
- setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
+ sum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"
- by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A)
+ by (simp add: exp_add exp_sum exp_diff mult_ac Gamma_complex_altdef A)
from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z
show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z"
by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def])
--- a/src/HOL/Analysis/Harmonic_Numbers.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Harmonic_Numbers.thy Mon Oct 17 14:37:32 2016 +0200
@@ -25,7 +25,7 @@
thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp
qed
-lemma setsum_Suc_diff':
+lemma sum_Suc_diff':
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
assumes "m \<le> n"
shows "(\<Sum>i = m..<n. f (Suc i) - f i) = f n - f m"
@@ -44,10 +44,10 @@
by (simp add: harm_def)
lemma harm_nonneg: "harm n \<ge> (0 :: 'a :: {real_normed_field,linordered_field})"
- unfolding harm_def by (intro setsum_nonneg) simp_all
+ unfolding harm_def by (intro sum_nonneg) simp_all
lemma harm_pos: "n > 0 \<Longrightarrow> harm n > (0 :: 'a :: {real_normed_field,linordered_field})"
- unfolding harm_def by (intro setsum_pos) simp_all
+ unfolding harm_def by (intro sum_pos) simp_all
lemma of_real_harm: "of_real (harm n) = harm n"
unfolding harm_def by simp
@@ -73,7 +73,7 @@
also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k=Suc 0..Suc n. inverse (of_nat k) :: real)"
unfolding harm_def[abs_def] by (subst convergent_Suc_iff) simp_all
also have "... \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k\<le>n. inverse (of_nat (Suc k)) :: real)"
- by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: atLeast0AtMost)
+ by (subst sum_shift_bounds_cl_Suc_ivl) (simp add: atLeast0AtMost)
also have "... \<longleftrightarrow> summable (\<lambda>n. inverse (of_nat n) :: real)"
by (subst summable_Suc_iff [symmetric]) (simp add: summable_iff_convergent')
also have "\<not>..." by (rule not_summable_harmonic)
@@ -126,7 +126,7 @@
"euler_mascheroni.sum_integral_diff_series n = harm (Suc n) - ln (of_nat (Suc n))"
proof -
have "harm (Suc n) = (\<Sum>k=0..n. inverse (of_nat k + 1) :: real)" unfolding harm_def
- unfolding One_nat_def by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: add_ac)
+ unfolding One_nat_def by (subst sum_shift_bounds_cl_Suc_ivl) (simp add: add_ac)
moreover have "((\<lambda>x. inverse (x + 1) :: real) has_integral ln (of_nat n + 1) - ln (0 + 1))
{0..of_nat n}"
by (intro fundamental_theorem_of_calculus)
@@ -201,16 +201,16 @@
fix n :: nat assume n: "n > 0"
have "(\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) =
(\<Sum>k<2*n. ((-1)^k + ?g k) / of_nat (Suc k)) - (\<Sum>k<2*n. ?g k / of_nat (Suc k))"
- by (simp add: setsum.distrib algebra_simps divide_inverse)
+ by (simp add: sum.distrib algebra_simps divide_inverse)
also have "(\<Sum>k<2*n. ((-1)^k + ?g k) / real_of_nat (Suc k)) = harm (2*n)"
- unfolding harm_altdef by (intro setsum.cong) (auto simp: field_simps)
+ unfolding harm_altdef by (intro sum.cong) (auto simp: field_simps)
also have "(\<Sum>k<2*n. ?g k / real_of_nat (Suc k)) = (\<Sum>k|k<2*n \<and> odd k. ?g k / of_nat (Suc k))"
- by (intro setsum.mono_neutral_right) auto
+ by (intro sum.mono_neutral_right) auto
also have "\<dots> = (\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k)))"
- by (intro setsum.cong) auto
+ by (intro sum.cong) auto
also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n"
unfolding harm_altdef
- by (intro setsum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE)
+ by (intro sum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE)
also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n
by (simp_all add: algebra_simps ln_mult)
finally show "?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))" ..
@@ -388,9 +388,9 @@
finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))"
by (simp add: inv_def field_simps)
also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
- unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf)
+ unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: sum.distrib sum_subtractf)
also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
- by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
+ by (subst atLeast0AtMost [symmetric], subst sum_Suc_diff) simp_all
finally show "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
by simp
@@ -422,9 +422,9 @@
finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))"
by (simp add: inv_def field_simps)
also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
- unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: setsum.distrib setsum_subtractf)
+ unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add: sum.distrib sum_subtractf)
also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
- by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
+ by (subst atLeast0AtMost [symmetric], subst sum_Suc_diff) simp_all
finally show "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
by simp
qed
--- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy Mon Oct 17 14:37:32 2016 +0200
@@ -176,11 +176,11 @@
unfolding forall_in_division[OF assms(2)]
by (metis antisym_conv assms content_pos_le content_subset division_ofD(2))
-lemma setsum_content_null:
+lemma sum_content_null:
assumes "content (cbox a b) = 0"
and "p tagged_division_of (cbox a b)"
- shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
-proof (rule setsum.neutral, rule)
+ shows "sum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
+proof (rule sum.neutral, rule)
fix y
assume y: "y \<in> p"
obtain x k where xk: "y = (x, k)"
@@ -199,12 +199,12 @@
lemma operative_content[intro]: "add.operative content"
by (force simp add: add.operative_def content_split[symmetric] content_eq_0_interior)
-lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> setsum content d = content (cbox a b)"
- by (metis operative_content setsum.operative_division)
+lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> sum content d = content (cbox a b)"
+ by (metis operative_content sum.operative_division)
lemma additive_content_tagged_division:
- "d tagged_division_of (cbox a b) \<Longrightarrow> setsum (\<lambda>(x,l). content l) d = content (cbox a b)"
- unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast
+ "d tagged_division_of (cbox a b) \<Longrightarrow> sum (\<lambda>(x,l). content l) d = content (cbox a b)"
+ unfolding sum.operative_tagged_division[OF operative_content, symmetric] by blast
lemma content_real_eq_0: "content {a .. b::real} = 0 \<longleftrightarrow> a \<ge> b"
by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0)
@@ -234,14 +234,14 @@
"(f has_integral y) (cbox a b) \<longleftrightarrow>
(\<forall>e>0. \<exists>d. gauge d \<and>
(\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
+ norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff)
lemma has_integral_real:
"(f has_integral y) {a .. b::real} \<longleftrightarrow>
(\<forall>e>0. \<exists>d. gauge d \<and>
(\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
+ norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
unfolding box_real[symmetric]
by (rule has_integral)
@@ -364,7 +364,7 @@
shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)"
using eventually_division_filter_tagged_division[of "cbox a b"]
additive_content_tagged_division[of _ a b]
- by (auto simp: has_integral_cbox split_beta' scaleR_setsum_left[symmetric]
+ by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric]
elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const])
lemma has_integral_const_real [intro]:
@@ -391,7 +391,7 @@
unfolding has_integral_cbox
using eventually_division_filter_tagged_division[of "cbox a b"]
by (subst tendsto_cong[where g="\<lambda>_. 0"])
- (auto elim!: eventually_mono intro!: setsum.neutral simp: tag_in_interval)
+ (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval)
{
presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
with assms lem show ?thesis
@@ -425,7 +425,7 @@
from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
by blast
have lem: "\<And>a b y f::'n\<Rightarrow>'a. (f has_integral y) (cbox a b) \<Longrightarrow> ((h \<circ> f) has_integral h y) (cbox a b)"
- unfolding has_integral_cbox by (drule tendsto) (simp add: setsum scaleR split_beta')
+ unfolding has_integral_cbox by (drule tendsto) (simp add: sum scaleR split_beta')
{
presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
then show ?thesis
@@ -531,7 +531,7 @@
((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)"
for f :: "'n \<Rightarrow> 'a" and g a b k l
unfolding has_integral_cbox
- by (simp add: split_beta' scaleR_add_right setsum.distrib[abs_def] tendsto_add)
+ by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add)
{
presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
then show ?thesis
@@ -666,10 +666,10 @@
shows "integral s (\<lambda>x. f x \<bullet> k) = integral s f \<bullet> k"
unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] ..
-lemma has_integral_setsum:
+lemma has_integral_sum:
assumes "finite t"
and "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
- shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum i t)) s"
+ shows "((\<lambda>x. sum (\<lambda>a. f a x) t) has_integral (sum i t)) s"
using assms(1) subset_refl[of t]
proof (induct rule: finite_subset_induct)
case empty
@@ -680,16 +680,16 @@
by (simp add: has_integral_add)
qed
-lemma integral_setsum:
+lemma integral_sum:
"\<lbrakk>finite t; \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow>
- integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t"
- by (auto intro: has_integral_setsum integrable_integral)
-
-lemma integrable_setsum:
- "\<lbrakk>finite t; \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s"
+ integral s (\<lambda>x. sum (\<lambda>a. f a x) t) = sum (\<lambda>a. integral s (f a)) t"
+ by (auto intro: has_integral_sum integrable_integral)
+
+lemma integrable_sum:
+ "\<lbrakk>finite t; \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> (\<lambda>x. sum (\<lambda>a. f a x) t) integrable_on s"
unfolding integrable_on_def
apply (drule bchoice)
- using has_integral_setsum[of t]
+ using has_integral_sum[of t]
apply auto
done
@@ -760,7 +760,7 @@
lemma has_integral_null [intro]: "content(cbox a b) = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
unfolding has_integral_cbox
using eventually_division_filter_tagged_division[of "cbox a b"]
- by (subst tendsto_cong[where g="\<lambda>_. 0"]) (auto elim: eventually_mono intro: setsum_content_null)
+ by (subst tendsto_cong[where g="\<lambda>_. 0"]) (auto elim: eventually_mono intro: sum_content_null)
lemma has_integral_null_real [intro]: "content {a .. b::real} = 0 \<Longrightarrow> (f has_integral 0) {a .. b}"
by (metis box_real(2) has_integral_null)
@@ -840,7 +840,7 @@
hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral ((y \<bullet> b) *\<^sub>R b)) A"
by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. (y \<bullet> b) *\<^sub>R b)) A"
- by (intro has_integral_setsum) (simp_all add: o_def)
+ by (intro has_integral_sum) (simp_all add: o_def)
thus "(f has_integral y) A" by (simp add: euclidean_representation)
qed
@@ -865,7 +865,7 @@
hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral (y b *\<^sub>R b)) A"
by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. y b *\<^sub>R b)) A"
- by (intro has_integral_setsum) (simp_all add: o_def)
+ by (intro has_integral_sum) (simp_all add: o_def)
thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation)
qed
@@ -885,7 +885,7 @@
have "integral A f = integral A (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b)"
by (simp add: euclidean_representation)
also from integrable have "\<dots> = (\<Sum>a\<in>Basis. integral A (\<lambda>x. (f x \<bullet> a) *\<^sub>R a))"
- by (subst integral_setsum) (simp_all add: o_def)
+ by (subst integral_sum) (simp_all add: o_def)
finally show ?thesis .
qed
@@ -944,7 +944,7 @@
from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]]
have dp: "\<And>i n. i\<le>n \<Longrightarrow> d i fine p n"
using p(2) unfolding fine_inters by auto
- have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
+ have "Cauchy (\<lambda>n. sum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
proof (rule CauchyI, goal_cases)
case (1 e)
then guess N unfolding real_arch_inverse[of e] .. note N=this
@@ -1124,9 +1124,9 @@
unfolding xk split_conv by auto
} note [simp] = this
have lem3: "\<And>g :: 'a set \<Rightarrow> 'a set. finite p \<Longrightarrow>
- setsum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} =
- setsum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)"
- by (rule setsum.mono_neutral_left) auto
+ sum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} =
+ sum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)"
+ by (rule sum.mono_neutral_left) auto
let ?M1 = "{(x, kk \<inter> {x. x\<bullet>k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
have d1_fine: "d1 fine ?M1"
by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm)
@@ -1218,10 +1218,10 @@
also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
unfolding lem3[OF p(3)]
- by (subst (1 2) setsum.reindex_nontrivial[OF p(3)])
+ by (subst (1 2) sum.reindex_nontrivial[OF p(3)])
(auto intro!: k eq0 tagged_division_split_left_inj_content[OF p(1)] tagged_division_split_right_inj_content[OF p(1)]
simp: cont_eq)+
- also note setsum.distrib[symmetric]
+ also note sum.distrib[symmetric]
also have "\<And>x. x \<in> p \<Longrightarrow>
(\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x +
(\<lambda>(x,ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x =
@@ -1236,7 +1236,7 @@
unfolding uv content_split[OF k,of u v c]
by auto
qed
- note setsum.cong [OF _ this]
+ note sum.cong [OF _ this]
finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x \<bullet> k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x \<bullet> k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)"
@@ -1258,7 +1258,7 @@
obtains d where "gauge d"
"\<forall>p1 p2. p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and>
p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2 \<longrightarrow>
- norm ((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e"
+ norm ((sum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + sum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e"
proof -
guess d using has_integralD[OF assms(1-2)] . note d=this
{ fix p1 p2
@@ -1287,9 +1287,9 @@
using e k by (auto simp: inner_simps inner_not_same_Basis)
have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
(\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))"
- using "*" by (blast intro: setsum.cong)
+ using "*" by (blast intro: sum.cong)
also have "\<dots> < e"
- apply (subst setsum.delta)
+ apply (subst sum.delta)
using e
apply auto
done
@@ -1308,7 +1308,7 @@
}
then have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) =
norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
- by (subst setsum.union_inter_neutral) (auto simp: p1 p2)
+ by (subst sum.union_inter_neutral) (auto simp: p1 p2)
also have "\<dots> < e"
by (rule k d(2) p12 fine_union p1 p2)+
finally have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" .
@@ -1426,20 +1426,20 @@
lemma dsum_bound:
assumes "p division_of (cbox a b)"
and "norm c \<le> e"
- shows "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)"
+ shows "norm (sum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)"
proof -
- have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = setsum content p"
- apply (rule setsum.cong)
+ have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = sum content p"
+ apply (rule sum.cong)
using assms
apply simp
apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4))
done
have e: "0 \<le> e"
using assms(2) norm_ge_zero order_trans by blast
- have "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
- using norm_setsum by blast
+ have "norm (sum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
+ using norm_sum by blast
also have "... \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)"
- by (simp add: setsum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono setsum_nonneg)
+ by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg)
also have "... \<le> e * content (cbox a b)"
apply (rule mult_left_mono [OF _ e])
apply (simp add: sumeq)
@@ -1451,7 +1451,7 @@
lemma rsum_bound:
assumes p: "p tagged_division_of (cbox a b)"
and "\<forall>x\<in>cbox a b. norm (f x) \<le> e"
- shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content (cbox a b)"
+ shows "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content (cbox a b)"
proof (cases "cbox a b = {}")
case True show ?thesis
using p unfolding True tagged_division_of_trivial by auto
@@ -1459,7 +1459,7 @@
case False
then have e: "e \<ge> 0"
by (meson ex_in_conv assms(2) norm_ge_zero order_trans)
- have setsum_le: "setsum (content \<circ> snd) p \<le> content (cbox a b)"
+ have sum_le: "sum (content \<circ> snd) p \<le> content (cbox a b)"
unfolding additive_content_tagged_division[OF p, symmetric] split_def
by (auto intro: eq_refl)
have con: "\<And>xk. xk \<in> p \<Longrightarrow> 0 \<le> content (snd xk)"
@@ -1468,15 +1468,15 @@
have norm: "\<And>xk. xk \<in> p \<Longrightarrow> norm (f (fst xk)) \<le> e"
unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms
by (metis prod.collapse subset_eq)
- have "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))"
- by (rule norm_setsum)
+ have "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))"
+ by (rule norm_sum)
also have "... \<le> e * content (cbox a b)"
unfolding split_def norm_scaleR
- apply (rule order_trans[OF setsum_mono])
+ apply (rule order_trans[OF sum_mono])
apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e])
apply (metis norm)
- unfolding setsum_distrib_right[symmetric]
- using con setsum_le
+ unfolding sum_distrib_right[symmetric]
+ using con sum_le
apply (auto simp: mult.commute intro: mult_left_mono [OF _ e])
done
finally show ?thesis .
@@ -1485,10 +1485,10 @@
lemma rsum_diff_bound:
assumes "p tagged_division_of (cbox a b)"
and "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e"
- shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le>
+ shows "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - sum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le>
e * content (cbox a b)"
apply (rule order_trans[OF _ rsum_bound[OF assms]])
- apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl)
+ apply (simp add: split_def scaleR_diff_right sum_subtractf eq_refl)
done
lemma has_integral_bound:
@@ -1534,9 +1534,9 @@
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
assumes "p tagged_division_of (cbox a b)"
and "\<forall>x\<in>cbox a b. (f x)\<bullet>i \<le> (g x)\<bullet>i"
- shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i"
-unfolding inner_setsum_left
-proof (rule setsum_mono, clarify)
+ shows "(sum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (sum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i"
+unfolding inner_sum_left
+proof (rule sum_mono, clarify)
fix a b
assume ab: "(a, b) \<in> p"
note tagged = tagged_division_ofD(2-4)[OF assms(1) ab]
@@ -1877,7 +1877,7 @@
also have "(\<Prod>j\<in>Basis. (b' 0 - a' 0) \<bullet> j) = 0"
using k *
by (intro setprod_zero bexI[OF _ k])
- (auto simp: b'_def a'_def inner_diff inner_setsum_left inner_not_same_Basis intro!: setsum.cong)
+ (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong)
also have "((\<lambda>d. \<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) \<longlongrightarrow> 0) (at_right 0) =
((\<lambda>d. content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d})) \<longlongrightarrow> 0) (at_right 0)"
proof (intro tendsto_cong eventually_at_rightI)
@@ -1952,7 +1952,7 @@
assume p: "p tagged_division_of (cbox a b) \<and> (\<lambda>x. ball x d) fine p"
have *: "(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) =
(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
apply cases
@@ -1989,7 +1989,7 @@
unfolding diff_0_right *
unfolding real_scaleR_def real_norm_def
apply (subst abs_of_nonneg)
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply rule
unfolding split_paired_all split_conv
apply (rule mult_nonneg_nonneg)
@@ -2005,14 +2005,14 @@
proof -
have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le>
(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
- apply (rule setsum_mono)
+ apply (rule sum_mono)
unfolding split_paired_all split_conv
apply (rule mult_right_le_one_le)
apply (drule p'(4))
apply (auto simp add:interval_doublesplit[OF k])
done
also have "\<dots> < e"
- proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases)
+ proof (subst sum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases)
case prems: (1 u v)
then have *: "content (cbox u v) = 0"
unfolding content_eq_0_interior by simp
@@ -2027,8 +2027,8 @@
by (blast intro: antisym)
next
have "(\<Sum>l\<in>snd ` p. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
- setsum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
- proof (subst (2) setsum.reindex_nontrivial)
+ sum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
+ proof (subst (2) sum.reindex_nontrivial)
fix x y assume "x \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}"
"x \<noteq> y" and eq: "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
then obtain x' y' where "(x', x) \<in> p" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> p" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}"
@@ -2041,7 +2041,7 @@
by (auto simp: eq)
then show "content (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
using p'(4)[OF \<open>(x', x) \<in> p\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int)
- qed (insert p'(1), auto intro!: setsum.mono_neutral_right)
+ qed (insert p'(1), auto intro!: sum.mono_neutral_right)
also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}. content l *\<^sub>R 1::real)"
by simp
also have "\<dots> \<le> 1 * content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
@@ -2144,15 +2144,15 @@
by (auto intro: tagged_division_finer[OF as(1) d(1)])
from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
have *: "\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> (0::real)"
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply safe
unfolding real_scaleR_def
apply (drule tagged_division_ofD(4)[OF q(1)])
apply (auto intro: mult_nonneg_nonneg)
done
have **: "finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow>
- (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t" for f g s t
- apply (rule setsum_le_included[of s t g snd f])
+ (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> sum f s \<le> sum g t" for f g s t
+ apply (rule sum_le_included[of s t g snd f])
prefer 4
apply safe
apply (erule_tac x=x in ballE)
@@ -2160,11 +2160,11 @@
apply (rule_tac x="(xa,x)" in bexI)
apply auto
done
- 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))) {..N+1}"
- unfolding real_norm_def setsum_distrib_left abs_of_nonneg[OF *] diff_0_right
+ have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> sum (\<lambda>i. (real i + 1) *
+ norm (sum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}"
+ unfolding real_norm_def sum_distrib_left abs_of_nonneg[OF *] diff_0_right
apply (rule order_trans)
- apply (rule norm_setsum)
+ apply (rule norm_sum)
apply (subst sum_sum_product)
prefer 3
proof (rule **, safe)
@@ -2229,8 +2229,8 @@
apply auto
done
qed (insert as, auto)
- also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}"
- proof (rule setsum_mono, goal_cases)
+ also have "\<dots> \<le> sum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}"
+ proof (rule sum_mono, goal_cases)
case (1 i)
then show ?case
apply (subst mult.commute, subst pos_le_divide_eq[symmetric])
@@ -2240,7 +2240,7 @@
done
qed
also have "\<dots> < e * inverse 2 * 2"
- unfolding divide_inverse setsum_distrib_left[symmetric]
+ unfolding divide_inverse sum_distrib_left[symmetric]
apply (rule mult_strict_left_mono)
unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric]
apply (subst geometric_sum)
@@ -2640,14 +2640,14 @@
lemma has_integral_factor_content:
"(f has_integral i) (cbox a b) \<longleftrightarrow>
(\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))"
+ norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))"
proof (cases "content (cbox a b) = 0")
case True
show ?thesis
unfolding has_integral_null_eq[OF True]
apply safe
apply (rule, rule, rule gauge_trivial, safe)
- unfolding setsum_content_null[OF True] True
+ unfolding sum_content_null[OF True] True
defer
apply (erule_tac x=1 in allE)
apply safe
@@ -2655,7 +2655,7 @@
apply (rule fine_division_exists[of _ a b])
apply assumption
apply (erule_tac x=p in allE)
- unfolding setsum_content_null[OF True]
+ unfolding sum_content_null[OF True]
apply auto
done
next
@@ -2696,7 +2696,7 @@
lemma has_integral_factor_content_real:
"(f has_integral i) {a .. b::real} \<longleftrightarrow>
(\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a .. b} ))"
+ norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a .. b} ))"
unfolding box_real[symmetric]
by (rule has_integral_factor_content)
@@ -2738,10 +2738,10 @@
show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)"
unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric]
unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\<lambda>x. x",symmetric]
- unfolding setsum_distrib_left
+ unfolding sum_distrib_left
defer
- unfolding setsum_subtractf[symmetric]
- proof (rule setsum_norm_le,safe)
+ unfolding sum_subtractf[symmetric]
+ proof (rule sum_norm_le,safe)
fix x k
assume "(x, k) \<in> p"
note xk = tagged_division_ofD(2-4)[OF as(1) this]
@@ -2809,7 +2809,7 @@
subsection \<open>Taylor series expansion\<close>
-lemma (in bounded_bilinear) setsum_prod_derivatives_has_vector_derivative:
+lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative:
assumes "p>0"
and f0: "Df 0 = f"
and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
@@ -2835,7 +2835,7 @@
prod (Df (Suc i) t) (Dg (p - Suc i) t))) =
(\<Sum>i\<le>(Suc p'). ?f i - ?f (Suc i))"
by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost)
- also note setsum_telescope
+ also note sum_telescope
finally
have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
prod (Df (Suc i) t) (Dg (p - Suc i) t)))
@@ -2884,7 +2884,7 @@
unfolding Dg_def
by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps)
let ?sum = "\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t"
- from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
+ from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
OF \<open>p > 0\<close> g0 Dg f0 Df]
have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
(?sum has_vector_derivative
@@ -2903,7 +2903,7 @@
then have "?sum b = f b"
using Suc_pred'[OF \<open>p > 0\<close>]
by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib
- cond_application_beta setsum.If_cases f0)
+ cond_application_beta sum.If_cases f0)
also
have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
proof safe
@@ -2914,7 +2914,7 @@
qed simp
from _ this
have "?sum a = (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a)"
- by (rule setsum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
+ by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
finally show c: ?case .
case 2 show ?case using c integral_unique by force
case 3 show ?case using c by force
@@ -3003,10 +3003,10 @@
unfolding euclidean_eq_iff[where 'a='a] using i by auto
have *: "Basis = insert i (Basis - {i})"
using i by auto
- have "norm (y - x) < e + setsum (\<lambda>i. 0) Basis"
+ have "norm (y - x) < e + sum (\<lambda>i. 0) Basis"
apply (rule le_less_trans[OF norm_le_l1])
apply (subst *)
- apply (subst setsum.insert)
+ apply (subst sum.insert)
prefer 3
apply (rule add_less_le_mono)
proof -
@@ -3436,9 +3436,9 @@
using inj(1) unfolding inj_on_def by fastforce
have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _")
using assms(7)
- apply (simp only: algebra_simps add_left_cancel scaleR_right.setsum)
- apply (subst setsum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
- apply (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4))
+ apply (simp only: algebra_simps add_left_cancel scaleR_right.sum)
+ apply (subst sum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
+ apply (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4))
done
also have "\<dots> = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r")
unfolding scaleR_diff_right scaleR_scaleR
@@ -3865,18 +3865,18 @@
have **: "\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2"
by arith
show ?case
- unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[symmetric] split_minus
- unfolding setsum_distrib_left
+ unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] sum_subtractf[symmetric] split_minus
+ unfolding sum_distrib_left
apply (subst(2) pA)
apply (subst pA)
- unfolding setsum.union_disjoint[OF pA(2-)]
+ unfolding sum.union_disjoint[OF pA(2-)]
proof (rule norm_triangle_le, rule **, goal_cases)
case 1
show ?case
apply (rule order_trans)
- apply (rule setsum_norm_le)
+ apply (rule sum_norm_le)
defer
- apply (subst setsum_divide_distrib)
+ apply (subst sum_divide_distrib)
apply (rule order_refl)
apply safe
apply (unfold not_le o_def split_conv fst_conv)
@@ -3930,12 +3930,12 @@
case 2
show ?case
apply (rule *)
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply rule
apply (unfold split_paired_all split_conv)
defer
- unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
- unfolding setsum_distrib_left[symmetric]
+ unfolding sum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
+ unfolding sum_distrib_left[symmetric]
apply (subst additive_tagged_division_1[OF _ as(1)])
apply (rule assms)
proof -
@@ -3948,12 +3948,12 @@
then show "0 \<le> e * ((Sup k) - (Inf k))"
unfolding uv using e by (auto simp add: field_simps)
next
- have *: "\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm (setsum f t) \<le> e \<Longrightarrow> norm (setsum f s) \<le> e"
+ have *: "\<And>s f t e. sum f s = sum f t \<Longrightarrow> norm (sum f t) \<le> e \<Longrightarrow> norm (sum f s) \<le> e"
by auto
show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x -
(f ((Sup k)) - f ((Inf k)))) \<le> e * (b - a) / 2"
apply (rule *[where t1="p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0}"])
- apply (rule setsum.mono_neutral_right[OF pA(2)])
+ apply (rule sum.mono_neutral_right[OF pA(2)])
defer
apply rule
unfolding split_paired_all split_conv o_def
@@ -3975,7 +3975,7 @@
have *: "p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0} =
{t. t\<in>p \<and> fst t = a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = b \<and> content(snd t) \<noteq> 0}"
by blast
- have **: "norm (setsum f s) \<le> e"
+ have **: "norm (sum f s) \<le> e"
if "\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y"
and "\<forall>x. x \<in> s \<longrightarrow> norm (f x) \<le> e"
and "e > 0"
@@ -3995,7 +3995,7 @@
case 2
show ?case
apply (subst *)
- apply (subst setsum.union_disjoint)
+ apply (subst sum.union_disjoint)
prefer 4
apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"])
apply (rule norm_triangle_le,rule add_mono)
@@ -4425,7 +4425,7 @@
then show ?thesis
unfolding ** box_real
apply -
- apply (subst setsum.insert)
+ apply (subst sum.insert)
apply (rule p')
unfolding split_conv
defer
@@ -5003,7 +5003,7 @@
show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm x \<le> B" and "i \<in> Basis"
using that and Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
unfolding c_def d_def
- by (auto simp add: field_simps setsum_negf)
+ by (auto simp add: field_simps sum_negf)
qed
have "ball 0 C \<subseteq> cbox c d"
apply (rule subsetI)
@@ -5014,7 +5014,7 @@
show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
using Basis_le_norm[OF i, of x] and x i
unfolding c_def d_def
- by (auto simp: setsum_negf)
+ by (auto simp: sum_negf)
qed
from C(2)[OF this] have "\<exists>y. (f has_integral y) (cbox a b)"
unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric]
@@ -5040,7 +5040,7 @@
then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
using Basis_le_norm[of i x]
unfolding c_def d_def
- by (auto simp add: field_simps setsum_negf)
+ by (auto simp add: field_simps sum_negf)
qed
have "ball 0 C \<subseteq> cbox c d"
apply (rule subsetI)
@@ -5051,7 +5051,7 @@
then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
using Basis_le_norm[of i x]
unfolding c_def d_def
- by (auto simp: setsum_negf)
+ by (auto simp: sum_negf)
qed
note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s]
note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]]
@@ -5425,7 +5425,7 @@
then show ?case
using Basis_le_norm[of i x] \<open>i\<in>Basis\<close>
using n N
- by (auto simp add: field_simps setsum_negf)
+ by (auto simp add: field_simps sum_negf)
qed
}
then show ?case
@@ -5482,7 +5482,7 @@
then show ?case
using Basis_le_norm[of i x] \<open>i \<in> Basis\<close>
using n
- by (auto simp add: field_simps setsum_negf)
+ by (auto simp add: field_simps sum_negf)
qed
qed
qed
@@ -5541,9 +5541,9 @@
and "0 \<le> (\<Sum>(x, k)\<in>p2. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)"
and "(\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R g x) \<ge> 0"
and "0 \<le> (\<Sum>(x, k)\<in>p1. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x)"
- unfolding setsum_subtractf[symmetric]
+ unfolding sum_subtractf[symmetric]
apply -
- apply (rule_tac[!] setsum_nonneg)
+ apply (rule_tac[!] sum_nonneg)
apply safe
unfolding real_scaleR_def right_diff_distrib[symmetric]
apply (rule_tac[!] mult_nonneg_nonneg)
@@ -5774,7 +5774,7 @@
assumes "finite t"
and "\<forall>s\<in>t. (f has_integral (i s)) s"
and "\<forall>s\<in>t. \<forall>s'\<in>t. s \<noteq> s' \<longrightarrow> negligible (s \<inter> s')"
- shows "(f has_integral (setsum i t)) (\<Union>t)"
+ shows "(f has_integral (sum i t)) (\<Union>t)"
proof -
note * = has_integral_restrict_univ[symmetric, of f]
have **: "negligible (\<Union>((\<lambda>(a,b). a \<inter> b) ` {(a,b). a \<in> t \<and> b \<in> {y. y \<in> t \<and> a \<noteq> y}}))"
@@ -5787,7 +5787,7 @@
apply auto
done
note assms(2)[unfolded *]
- note has_integral_setsum[OF assms(1) this]
+ note has_integral_sum[OF assms(1) this]
then show ?thesis
unfolding *
apply -
@@ -5807,12 +5807,12 @@
unfolding if_P[OF True]
apply (rule trans)
defer
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
apply (subst *)
apply assumption
apply (rule refl)
- unfolding setsum.delta[OF assms(1)]
+ unfolding sum.delta[OF assms(1)]
using s
apply auto
done
@@ -5827,7 +5827,7 @@
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "d division_of s"
and "\<forall>k\<in>d. (f has_integral (i k)) k"
- shows "(f has_integral (setsum i d)) s"
+ shows "(f has_integral (sum i d)) s"
proof -
note d = division_ofD[OF assms(1)]
show ?thesis
@@ -5854,7 +5854,7 @@
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "d division_of s"
and "\<forall>k\<in>d. f integrable_on k"
- shows "integral s f = setsum (\<lambda>i. integral i f) d"
+ shows "integral s f = sum (\<lambda>i. integral i f) d"
apply (rule integral_unique)
apply (rule has_integral_combine_division[OF assms(1)])
using assms(2)
@@ -5867,7 +5867,7 @@
assumes "f integrable_on s"
and "d division_of k"
and "k \<subseteq> s"
- shows "(f has_integral (setsum (\<lambda>i. integral i f) d)) k"
+ shows "(f has_integral (sum (\<lambda>i. integral i f) d)) k"
apply (rule has_integral_combine_division[OF assms(2)])
apply safe
unfolding has_integral_integral[symmetric]
@@ -5887,7 +5887,7 @@
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on s"
and "d division_of s"
- shows "integral s f = setsum (\<lambda>i. integral i f) d"
+ shows "integral s f = sum (\<lambda>i. integral i f) d"
apply (rule integral_unique)
apply (rule has_integral_combine_division_topdown)
using assms
@@ -5941,17 +5941,17 @@
apply auto
done
also have "(\<Sum>k\<in>snd`p. integral k f) = (\<Sum>(x, k)\<in>p. integral k f)"
- by (intro setsum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null)
+ by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null)
(simp add: content_eq_0_interior)
finally show ?thesis
- using assms by (auto simp add: has_integral_iff intro!: setsum.cong)
+ using assms by (auto simp add: has_integral_iff intro!: sum.cong)
qed
lemma integral_combine_tagged_division_bottomup:
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "p tagged_division_of (cbox a b)"
and "\<forall>(x,k)\<in>p. f integrable_on k"
- shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
+ shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p"
apply (rule integral_unique)
apply (rule has_integral_combine_tagged_division[OF assms(1)])
using assms(2)
@@ -5962,7 +5962,7 @@
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on cbox a b"
and "p tagged_division_of (cbox a b)"
- shows "(f has_integral (setsum (\<lambda>(x,k). integral k f) p)) (cbox a b)"
+ shows "(f has_integral (sum (\<lambda>(x,k). integral k f) p)) (cbox a b)"
apply (rule has_integral_combine_tagged_division[OF assms(2)])
apply safe
proof goal_cases
@@ -5976,7 +5976,7 @@
fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
assumes "f integrable_on cbox a b"
and "p tagged_division_of (cbox a b)"
- shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
+ shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p"
apply (rule integral_unique)
apply (rule has_integral_combine_tagged_division_topdown)
using assms
@@ -5992,9 +5992,9 @@
and "e > 0"
and "gauge d"
and "(\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)"
+ norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)"
and p: "p tagged_partial_division_of (cbox a b)" "d fine p"
- shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x - integral k f) p) \<le> e"
+ shows "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x - integral k f) p) \<le> e"
(is "?x \<le> e")
proof -
{ presume "\<And>k. 0<k \<Longrightarrow> ?x \<le> e + k" then show ?thesis by (blast intro: field_le_epsilon) }
@@ -6012,7 +6012,7 @@
using q' unfolding r_def by auto
have "\<forall>i\<in>r. \<exists>p. p tagged_division_of i \<and> d fine p \<and>
- norm (setsum (\<lambda>(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)"
+ norm (sum (\<lambda>(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)"
apply safe
proof goal_cases
case (1 i)
@@ -6100,7 +6100,7 @@
then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>\<Union>(qq ` r). content k *\<^sub>R f x) -
integral (cbox a b) f) < e"
- apply (subst setsum.union_inter_neutral[symmetric])
+ apply (subst sum.union_inter_neutral[symmetric])
apply (rule p')
prefer 3
apply assumption
@@ -6124,9 +6124,9 @@
unfolding uv content_eq_0_interior[symmetric] by auto
qed auto
- then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x))
+ then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + sum (sum (\<lambda>(x, k). content k *\<^sub>R f x))
(qq ` r) - integral (cbox a b) f) < e"
- apply (subst (asm) setsum.Union_comp)
+ apply (subst (asm) sum.Union_comp)
prefer 2
unfolding split_paired_all split_conv image_iff
apply (erule bexE)+
@@ -6150,11 +6150,11 @@
by auto
qed (insert qq, auto)
- then have **: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x) \<circ> qq) r -
+ then have **: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + sum (sum (\<lambda>(x, k). content k *\<^sub>R f x) \<circ> qq) r -
integral (cbox a b) f) < e"
- apply (subst (asm) setsum.reindex_nontrivial)
+ apply (subst (asm) sum.reindex_nontrivial)
apply fact
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply rule
unfolding split_paired_all split_conv
defer
@@ -6180,28 +6180,28 @@
qed
have "?x = norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
- unfolding split_def setsum_subtractf ..
+ unfolding split_def sum_subtractf ..
also have "\<dots> \<le> e + k"
- apply (rule *[OF **, where ir1="setsum (\<lambda>k. integral k f) r"])
+ apply (rule *[OF **, where ir1="sum (\<lambda>k. integral k f) r"])
proof goal_cases
case 1
have *: "k * real (card r) / (1 + real (card r)) < k"
using k by (auto simp add: field_simps)
show ?case
- apply (rule le_less_trans[of _ "setsum (\<lambda>x. k / (real (card r) + 1)) r"])
- unfolding setsum_subtractf[symmetric]
- apply (rule setsum_norm_le)
+ apply (rule le_less_trans[of _ "sum (\<lambda>x. k / (real (card r) + 1)) r"])
+ unfolding sum_subtractf[symmetric]
+ apply (rule sum_norm_le)
apply rule
apply (drule qq)
defer
- unfolding divide_inverse setsum_distrib_right[symmetric]
+ unfolding divide_inverse sum_distrib_right[symmetric]
unfolding divide_inverse[symmetric]
using * apply (auto simp add: field_simps)
done
next
case 2
have *: "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)"
- apply (subst setsum.reindex_nontrivial)
+ apply (subst sum.reindex_nontrivial)
apply fact
unfolding split_paired_all snd_conv split_def o_def
proof -
@@ -6221,7 +6221,7 @@
from q(1) have **: "snd ` p \<union> q = q" by auto
show ?case
unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def
- using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\<lambda>k. integral k f", symmetric]
+ using ** q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p" "\<lambda>k. integral k f", symmetric]
by simp
qed
finally show "?x \<le> e + k" .
@@ -6233,12 +6233,12 @@
and "e > 0"
and "gauge d"
and "\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
- norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e"
+ norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e"
and "p tagged_partial_division_of (cbox a b)"
and "d fine p"
- shows "setsum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
+ shows "sum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
unfolding split_def
- apply (rule setsum_norm_allsubsets_bound)
+ apply (rule sum_norm_allsubsets_bound)
defer
apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)])
apply safe
@@ -6260,7 +6260,7 @@
and "e > 0"
obtains d where "gauge d"
and "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d fine p \<longrightarrow>
- setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
+ sum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
proof -
have *: "e / (2 * (real DIM('n) + 1)) > 0" using assms(2) by simp
from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this]
@@ -6287,7 +6287,7 @@
lemma sum_gp_basic:
fixes x :: "'a::ring_1"
- shows "(1 - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
+ shows "(1 - x) * sum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
proof -
define y where "y = 1 - x"
have "y * (\<Sum>i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n"
@@ -6298,7 +6298,7 @@
lemma sum_gp_multiplied:
assumes mn: "m \<le> n"
- shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
+ shows "((1::'a::{field}) - x) * sum (op ^ x) {m..n} = x^m - x^ Suc n"
(is "?lhs = ?rhs")
proof -
let ?S = "{0..(n - m)}"
@@ -6314,8 +6314,8 @@
done
have th: "op ^ x \<circ> op + m = (\<lambda>i. x^m * x^i)"
by (rule ext) (simp add: power_add power_mult)
- from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_distrib_left[symmetric]]
- have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})"
+ from sum.reindex[OF i, of "op ^ x", unfolded f th sum_distrib_left[symmetric]]
+ have "?lhs = x^m * ((1 - x) * sum (op ^ x) {0..n - m})"
by simp
then show ?thesis
unfolding sum_gp_basic
@@ -6324,7 +6324,7 @@
qed
lemma sum_gp:
- "setsum (op ^ (x::'a::{field})) {m .. n} =
+ "sum (op ^ (x::'a::{field})) {m .. n} =
(if n < m then 0
else if x = 1 then of_nat ((n + 1) - m)
else (x^ m - x^ (Suc n)) / (1 - x))"
@@ -6357,7 +6357,7 @@
qed
lemma sum_gp_offset:
- "setsum (op ^ (x::'a::{field})) {m .. m+n} =
+ "sum (op ^ (x::'a::{field})) {m .. m+n} =
(if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
unfolding sum_gp[of x m "m + n"] power_Suc
by (simp add: field_simps power_add)
@@ -6497,12 +6497,12 @@
case 1
show ?case
apply (rule order_trans[of _ "\<Sum>(x, k)\<in>p. content k * (e / (4 * content (cbox a b)))"])
- unfolding setsum_subtractf[symmetric]
+ unfolding sum_subtractf[symmetric]
apply (rule order_trans)
- apply (rule norm_setsum)
- apply (rule setsum_mono)
+ apply (rule norm_sum)
+ apply (rule sum_mono)
unfolding split_paired_all split_conv
- unfolding split_def setsum_distrib_right[symmetric] scaleR_diff_right[symmetric]
+ unfolding split_def sum_distrib_right[symmetric] scaleR_diff_right[symmetric]
unfolding additive_content_tagged_division[OF p(1), unfolded split_def]
proof -
fix x k
@@ -6524,22 +6524,22 @@
show ?case
apply (rule le_less_trans[of _ "norm (\<Sum>j = 0..s.
\<Sum>(x, k)\<in>{xk\<in>p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))"])
- apply (subst setsum_group)
+ apply (subst sum_group)
apply fact
apply (rule finite_atLeastAtMost)
defer
apply (subst split_def)+
- unfolding setsum_subtractf
+ unfolding sum_subtractf
apply rule
proof -
show "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> p.
m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e / 2"
- apply (rule le_less_trans[of _ "setsum (\<lambda>i. e / 2^(i+2)) {0..s}"])
- apply (rule setsum_norm_le)
+ apply (rule le_less_trans[of _ "sum (\<lambda>i. e / 2^(i+2)) {0..s}"])
+ apply (rule sum_norm_le)
proof
show "(\<Sum>i = 0..s. e / 2 ^ (i + 2)) < e / 2"
unfolding power_add divide_inverse inverse_mult_distrib
- unfolding setsum_distrib_left[symmetric] setsum_distrib_right[symmetric]
+ unfolding sum_distrib_left[symmetric] sum_distrib_right[symmetric]
unfolding power_inverse [symmetric] sum_gp
apply(rule mult_strict_left_mono[OF _ e])
unfolding power2_eq_square
@@ -6550,10 +6550,10 @@
show "norm (\<Sum>(x, k)\<in>{xk \<in> p. m (fst xk) = t}. content k *\<^sub>R f (m x) x -
integral k (f (m x))) \<le> e / 2 ^ (t + 2)"
apply (rule order_trans
- [of _ "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \<in> p. m (fst xk) = t})"])
+ [of _ "norm (sum (\<lambda>(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \<in> p. m (fst xk) = t})"])
apply (rule eq_refl)
apply (rule arg_cong[where f=norm])
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
defer
apply (rule henstock_lemma_part1)
@@ -6587,7 +6587,7 @@
apply (rule comb[of r])
apply (rule comb[of s])
apply (rule i'[unfolded real_inner_1_right])
- apply (rule_tac[1-2] setsum_mono)
+ apply (rule_tac[1-2] sum_mono)
unfolding split_paired_all split_conv
apply (rule_tac[1-2] integral_le[OF ])
proof safe
@@ -7013,7 +7013,7 @@
defer
apply (rule d1(2)[OF conjI[OF p(1)]])
defer
- apply (rule setsum_norm_le)
+ apply (rule sum_norm_le)
proof safe
fix x k
assume "(x, k) \<in> p"
--- a/src/HOL/Analysis/Homeomorphism.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Homeomorphism.thy Mon Oct 17 14:37:32 2016 +0200
@@ -999,22 +999,22 @@
proof
show "linear f"
unfolding f_def
- by (intro linear_compose_setsum linearI ballI) (auto simp: algebra_simps)
+ by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps)
show "linear g"
unfolding g_def
- by (intro linear_compose_setsum linearI ballI) (auto simp: algebra_simps)
+ by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps)
have *: "(\<Sum>a \<in> Basis. a \<bullet> basf b * (x \<bullet> basg a)) = x \<bullet> b" if "b \<in> Basis" for x b
using sbf that by auto
show gf: "g (f x) = x" for x
apply (rule euclidean_eqI)
- apply (simp add: f_def g_def inner_setsum_left scaleR_setsum_left algebra_simps)
- apply (simp add: Groups_Big.setsum_distrib_left [symmetric] *)
+ apply (simp add: f_def g_def inner_sum_left scaleR_sum_left algebra_simps)
+ apply (simp add: Groups_Big.sum_distrib_left [symmetric] *)
done
show "basf(0,1) \<in> Basis"
using b01 sbf by auto
then show "f(x,0) \<bullet> basf(0,1) = 0" for x
- apply (simp add: f_def inner_setsum_left)
- apply (rule comm_monoid_add_class.setsum.neutral)
+ apply (simp add: f_def inner_sum_left)
+ apply (rule comm_monoid_add_class.sum.neutral)
using b01 inner_not_same_Basis by fastforce
qed
qed
--- a/src/HOL/Analysis/Inner_Product.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Inner_Product.thy Mon Oct 17 14:37:32 2016 +0200
@@ -46,7 +46,7 @@
lemma inner_diff_left: "inner (x - y) z = inner x z - inner y z"
using inner_add_left [of x "- y" z] by simp
-lemma inner_setsum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
+lemma inner_sum_left: "inner (\<Sum>x\<in>A. f x) y = (\<Sum>x\<in>A. inner (f x) y)"
by (cases "finite A", induct set: finite, simp_all add: inner_add_left)
text \<open>Transfer distributivity rules to right argument.\<close>
@@ -66,8 +66,8 @@
lemma inner_diff_right: "inner x (y - z) = inner x y - inner x z"
using inner_diff_left [of y z x] by (simp only: inner_commute)
-lemma inner_setsum_right: "inner x (\<Sum>y\<in>A. f y) = (\<Sum>y\<in>A. inner x (f y))"
- using inner_setsum_left [of f A x] by (simp only: inner_commute)
+lemma inner_sum_right: "inner x (\<Sum>y\<in>A. f y) = (\<Sum>y\<in>A. inner x (f y))"
+ using inner_sum_left [of f A x] by (simp only: inner_commute)
lemmas inner_add [algebra_simps] = inner_add_left inner_add_right
lemmas inner_diff [algebra_simps] = inner_diff_left inner_diff_right
--- a/src/HOL/Analysis/Integral_Test.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Integral_Test.thy Mon Oct 17 14:37:32 2016 +0200
@@ -43,11 +43,11 @@
with Suc show ?case by simp
qed simp_all
also have "... \<le> (\<Sum>k<n. integral {of_nat k..of_nat (Suc k)} (\<lambda>_::real. f (of_nat k)))"
- by (intro setsum_mono integral_le int) (auto intro: dec)
+ by (intro sum_mono integral_le int) (auto intro: dec)
also have "... = (\<Sum>k<n. f (of_nat k))" by simp
also have "\<dots> - (\<Sum>k\<le>n. f (of_nat k)) = -(\<Sum>k\<in>{..n} - {..<n}. f (of_nat k))"
- by (subst setsum_diff) auto
- also have "\<dots> \<le> 0" by (auto intro!: setsum_nonneg nonneg)
+ by (subst sum_diff) auto
+ also have "\<dots> \<le> 0" by (auto intro!: sum_nonneg nonneg)
finally show "sum_integral_diff_series n \<ge> 0" by simp
qed
--- a/src/HOL/Analysis/L2_Norm.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/L2_Norm.thy Mon Oct 17 14:37:32 2016 +0200
@@ -28,10 +28,10 @@
lemma setL2_insert [simp]:
"\<lbrakk>finite F; a \<notin> F\<rbrakk> \<Longrightarrow>
setL2 f (insert a F) = sqrt ((f a)\<^sup>2 + (setL2 f F)\<^sup>2)"
- unfolding setL2_def by (simp add: setsum_nonneg)
+ unfolding setL2_def by (simp add: sum_nonneg)
lemma setL2_nonneg [simp]: "0 \<le> setL2 f A"
- unfolding setL2_def by (simp add: setsum_nonneg)
+ unfolding setL2_def by (simp add: sum_nonneg)
lemma setL2_0': "\<forall>a\<in>A. f a = 0 \<Longrightarrow> setL2 f A = 0"
unfolding setL2_def by simp
@@ -44,7 +44,7 @@
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
shows "setL2 f K \<le> setL2 g K"
unfolding setL2_def
- by (simp add: setsum_nonneg setsum_mono power_mono assms)
+ by (simp add: sum_nonneg sum_mono power_mono assms)
lemma setL2_strict_mono:
assumes "finite K" and "K \<noteq> {}"
@@ -52,27 +52,27 @@
assumes "\<And>i. i \<in> K \<Longrightarrow> 0 \<le> f i"
shows "setL2 f K < setL2 g K"
unfolding setL2_def
- by (simp add: setsum_strict_mono power_strict_mono assms)
+ by (simp add: sum_strict_mono power_strict_mono assms)
lemma setL2_right_distrib:
"0 \<le> r \<Longrightarrow> r * setL2 f A = setL2 (\<lambda>x. r * f x) A"
unfolding setL2_def
apply (simp add: power_mult_distrib)
- apply (simp add: setsum_distrib_left [symmetric])
- apply (simp add: real_sqrt_mult setsum_nonneg)
+ apply (simp add: sum_distrib_left [symmetric])
+ apply (simp add: real_sqrt_mult sum_nonneg)
done
lemma setL2_left_distrib:
"0 \<le> r \<Longrightarrow> setL2 f A * r = setL2 (\<lambda>x. f x * r) A"
unfolding setL2_def
apply (simp add: power_mult_distrib)
- apply (simp add: setsum_distrib_right [symmetric])
- apply (simp add: real_sqrt_mult setsum_nonneg)
+ apply (simp add: sum_distrib_right [symmetric])
+ apply (simp add: real_sqrt_mult sum_nonneg)
done
lemma setL2_eq_0_iff: "finite A \<Longrightarrow> setL2 f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
unfolding setL2_def
- by (simp add: setsum_nonneg setsum_nonneg_eq_0_iff)
+ by (simp add: sum_nonneg sum_nonneg_eq_0_iff)
lemma setL2_triangle_ineq:
shows "setL2 (\<lambda>i. f i + g i) A \<le> setL2 f A + setL2 g A"
@@ -106,8 +106,8 @@
apply simp
done
-lemma setL2_le_setsum [rule_format]:
- "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> setsum f A"
+lemma setL2_le_sum [rule_format]:
+ "(\<forall>i\<in>A. 0 \<le> f i) \<longrightarrow> setL2 f A \<le> sum f A"
apply (cases "finite A")
apply (induct set: finite)
apply simp
@@ -124,7 +124,7 @@
apply simp
done
-lemma setL2_le_setsum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
+lemma setL2_le_sum_abs: "setL2 f A \<le> (\<Sum>i\<in>A. \<bar>f i\<bar>)"
apply (cases "finite A")
apply (induct set: finite)
apply simp
@@ -155,7 +155,7 @@
apply (rule order_trans)
apply (rule power_mono)
apply (erule add_left_mono)
- apply (simp add: setsum_nonneg)
+ apply (simp add: sum_nonneg)
apply (simp add: power2_sum)
apply (simp add: power_mult_distrib)
apply (simp add: distrib_left distrib_right)
@@ -166,6 +166,6 @@
lemma member_le_setL2: "\<lbrakk>finite A; i \<in> A\<rbrakk> \<Longrightarrow> f i \<le> setL2 f A"
unfolding setL2_def
- by (auto intro!: member_le_setsum real_le_rsqrt)
+ by (auto intro!: member_le_sum real_le_rsqrt)
end
--- a/src/HOL/Analysis/Lebesgue_Measure.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Lebesgue_Measure.thy Mon Oct 17 14:37:32 2016 +0200
@@ -82,7 +82,7 @@
by fastforce
have "(\<Sum>i\<in>S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\<Sum>i\<in>S - {m}. F (r i) - F (l i))"
- using m psubset by (intro setsum.remove) auto
+ using m psubset by (intro sum.remove) auto
also have "(\<Sum>i\<in>S - {m}. F (r i) - F (l i)) \<le> F b - F (r m)"
proof (intro psubset.IH)
show "S - {m} \<subset> S"
@@ -144,7 +144,7 @@
by (intro psubset) auto
also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
using psubset.prems
- by (intro setsum_mono2 psubset) (auto intro: less_imp_le)
+ by (intro sum_mono2 psubset) (auto intro: less_imp_le)
finally show ?thesis .
next
assume "\<not> ?R"
@@ -155,13 +155,13 @@
have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
- by (intro setsum_mono2) (auto intro: less_imp_le)
+ by (intro sum_mono2) (auto intro: less_imp_le)
also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
(\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
using psubset(1) psubset.prems(1) j
- apply (subst setsum.union_disjoint)
+ apply (subst sum.union_disjoint)
apply simp_all
- apply (subst setsum.union_disjoint)
+ apply (subst sum.union_disjoint)
apply auto
apply (metis less_le_not_le)
done
@@ -245,7 +245,7 @@
apply arith
by (rule deltai_gt0)
also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply simp
apply (rule order_trans)
apply (rule less_imp_le)
@@ -253,11 +253,11 @@
by auto
also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
(epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
- by (subst setsum.distrib) (simp add: field_simps setsum_distrib_left)
+ by (subst sum.distrib) (simp add: field_simps sum_distrib_left)
also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
apply (rule add_left_mono)
apply (rule mult_left_mono)
- apply (rule setsum_mono2)
+ apply (rule sum_mono2)
using egt0 apply auto
by (frule Sbound, auto)
also have "... \<le> ?t + (epsilon / 2)"
@@ -281,11 +281,11 @@
also have "... = (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon"
by auto
also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
- using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
+ using finS Sbound Sprop by (auto intro!: add_right_mono sum_mono3)
finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
- using egt0 by (simp add: ennreal_plus[symmetric] setsum_nonneg del: ennreal_plus)
+ using egt0 by (simp add: ennreal_plus[symmetric] sum_nonneg del: ennreal_plus)
then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
- by (rule order_trans) (auto intro!: add_mono setsum_le_suminf simp del: setsum_ennreal)
+ by (rule order_trans) (auto intro!: add_mono sum_le_suminf simp del: sum_ennreal)
qed
moreover have "(\<Sum>i. ennreal (F (r i) - F (l i))) \<le> ennreal (F b - F a)"
using \<open>a \<le> b\<close> by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
@@ -637,7 +637,7 @@
fixes t :: "'a::euclidean_space"
shows "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c *\<^sub>R x)) (\<lambda>_. \<bar>c\<bar>^DIM('a))"
using lborel_affine_euclidean[where c="\<lambda>_::'a. c" and t=t]
- unfolding scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric] euclidean_representation setprod_constant by simp
+ unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation setprod_constant by simp
lemma lborel_real_affine:
"c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs c))"
@@ -722,7 +722,7 @@
next
assume "x \<noteq> 0" then show ?thesis
using lebesgue_affine_measurable[of "\<lambda>_. x" 0]
- unfolding scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric] euclidean_representation
+ unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation
by (auto simp add: ac_simps)
qed
@@ -747,7 +747,7 @@
using bij_vimage_eq_inv_image[OF \<open>bij ?T'\<close>, of S] by auto
have trans_eq_T: "(\<lambda>x. \<delta> + (\<Sum>j\<in>Basis. (m * (x \<bullet> j)) *\<^sub>R j)) = T m \<delta>" for m \<delta>
- unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric]
+ unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_sum_right[symmetric]
by (auto simp add: euclidean_representation ac_simps)
have T[measurable]: "T r d \<in> lebesgue \<rightarrow>\<^sub>M lebesgue" for r d
--- a/src/HOL/Analysis/Linear_Algebra.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Linear_Algebra.thy Mon Oct 17 14:37:32 2016 +0200
@@ -141,9 +141,9 @@
lemma linear_uminus: "linear uminus"
by (simp add: linear_iff)
-lemma linear_compose_setsum:
+lemma linear_compose_sum:
assumes lS: "\<forall>a \<in> S. linear (f a)"
- shows "linear (\<lambda>x. setsum (\<lambda>a. f a x) S)"
+ shows "linear (\<lambda>x. sum (\<lambda>a. f a x) S)"
proof (cases "finite S")
case True
then show ?thesis
@@ -173,9 +173,9 @@
lemma linear_diff: "linear f \<Longrightarrow> f (x - y) = f x - f y"
using linear_add [of f x "- y"] by (simp add: linear_neg)
-lemma linear_setsum:
+lemma linear_sum:
assumes f: "linear f"
- shows "f (setsum g S) = setsum (f \<circ> g) S"
+ shows "f (sum g S) = sum (f \<circ> g) S"
proof (cases "finite S")
case True
then show ?thesis
@@ -186,10 +186,10 @@
by (simp add: linear_0 [OF f])
qed
-lemma linear_setsum_mul:
+lemma linear_sum_mul:
assumes lin: "linear f"
- shows "f (setsum (\<lambda>i. c i *\<^sub>R v i) S) = setsum (\<lambda>i. c i *\<^sub>R f (v i)) S"
- using linear_setsum[OF lin, of "\<lambda>i. c i *\<^sub>R v i" , unfolded o_def] linear_cmul[OF lin]
+ shows "f (sum (\<lambda>i. c i *\<^sub>R v i) S) = sum (\<lambda>i. c i *\<^sub>R f (v i)) S"
+ using linear_sum[OF lin, of "\<lambda>i. c i *\<^sub>R v i" , unfolded o_def] linear_cmul[OF lin]
by simp
lemma linear_injective_0:
@@ -250,10 +250,10 @@
lemma subspace_diff: "subspace S \<Longrightarrow> x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x - y \<in> S"
using subspace_add [of S x "- y"] by (simp add: subspace_neg)
-lemma (in real_vector) subspace_setsum:
+lemma (in real_vector) subspace_sum:
assumes sA: "subspace A"
and f: "\<And>x. x \<in> B \<Longrightarrow> f x \<in> A"
- shows "setsum f B \<in> A"
+ shows "sum f B \<in> A"
proof (cases "finite B")
case True
then show ?thesis
@@ -460,8 +460,8 @@
lemma span_diff: "x \<in> span S \<Longrightarrow> y \<in> span S \<Longrightarrow> x - y \<in> span S"
by (metis subspace_span subspace_diff)
-lemma (in real_vector) span_setsum: "(\<And>x. x \<in> A \<Longrightarrow> f x \<in> span S) \<Longrightarrow> setsum f A \<in> span S"
- by (rule subspace_setsum [OF subspace_span])
+lemma (in real_vector) span_sum: "(\<And>x. x \<in> A \<Longrightarrow> f x \<in> span S) \<Longrightarrow> sum f A \<in> span S"
+ by (rule subspace_sum [OF subspace_span])
lemma span_add_eq: "x \<in> span S \<Longrightarrow> x + y \<in> span S \<longleftrightarrow> y \<in> span S"
by (metis add_minus_cancel scaleR_minus1_left subspace_def subspace_span)
@@ -619,16 +619,16 @@
text \<open>An explicit expansion is sometimes needed.\<close>
lemma span_explicit:
- "span P = {y. \<exists>S u. finite S \<and> S \<subseteq> P \<and> setsum (\<lambda>v. u v *\<^sub>R v) S = y}"
+ "span P = {y. \<exists>S u. finite S \<and> S \<subseteq> P \<and> sum (\<lambda>v. u v *\<^sub>R v) S = y}"
(is "_ = ?E" is "_ = {y. ?h y}" is "_ = {y. \<exists>S u. ?Q S u y}")
proof -
{
fix x
assume "?h x"
- then obtain S u where "finite S" and "S \<subseteq> P" and "setsum (\<lambda>v. u v *\<^sub>R v) S = x"
+ then obtain S u where "finite S" and "S \<subseteq> P" and "sum (\<lambda>v. u v *\<^sub>R v) S = x"
by blast
then have "x \<in> span P"
- by (auto intro: span_setsum span_mul span_superset)
+ by (auto intro: span_sum span_mul span_superset)
}
moreover
have "\<forall>x \<in> span P. ?h x"
@@ -640,7 +640,7 @@
assume x: "x \<in> P"
assume hy: "?h y"
from hy obtain S u where fS: "finite S" and SP: "S\<subseteq>P"
- and u: "setsum (\<lambda>v. u v *\<^sub>R v) S = y" by blast
+ and u: "sum (\<lambda>v. u v *\<^sub>R v) S = y" by blast
let ?S = "insert x S"
let ?u = "\<lambda>y. if y = x then (if x \<in> S then u y + c else c) else u y"
from fS SP x have th0: "finite (insert x S)" "insert x S \<subseteq> P"
@@ -648,19 +648,19 @@
have "?Q ?S ?u (c*\<^sub>R x + y)"
proof cases
assume xS: "x \<in> S"
- have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = (\<Sum>v\<in>S - {x}. u v *\<^sub>R v) + (u x + c) *\<^sub>R x"
- using xS by (simp add: setsum.remove [OF fS xS] insert_absorb)
+ have "sum (\<lambda>v. ?u v *\<^sub>R v) ?S = (\<Sum>v\<in>S - {x}. u v *\<^sub>R v) + (u x + c) *\<^sub>R x"
+ using xS by (simp add: sum.remove [OF fS xS] insert_absorb)
also have "\<dots> = (\<Sum>v\<in>S. u v *\<^sub>R v) + c *\<^sub>R x"
- by (simp add: setsum.remove [OF fS xS] algebra_simps)
+ by (simp add: sum.remove [OF fS xS] algebra_simps)
also have "\<dots> = c*\<^sub>R x + y"
by (simp add: add.commute u)
- finally have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = c*\<^sub>R x + y" .
+ finally have "sum (\<lambda>v. ?u v *\<^sub>R v) ?S = c*\<^sub>R x + y" .
then show ?thesis using th0 by blast
next
assume xS: "x \<notin> S"
have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *\<^sub>R v) = y"
unfolding u[symmetric]
- apply (rule setsum.cong)
+ apply (rule sum.cong)
using xS
apply auto
done
@@ -674,13 +674,13 @@
qed
lemma dependent_explicit:
- "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>v\<in>S. u v \<noteq> 0 \<and> setsum (\<lambda>v. u v *\<^sub>R v) S = 0))"
+ "dependent P \<longleftrightarrow> (\<exists>S u. finite S \<and> S \<subseteq> P \<and> (\<exists>v\<in>S. u v \<noteq> 0 \<and> sum (\<lambda>v. u v *\<^sub>R v) S = 0))"
(is "?lhs = ?rhs")
proof -
{
assume dP: "dependent P"
then obtain a S u where aP: "a \<in> P" and fS: "finite S"
- and SP: "S \<subseteq> P - {a}" and ua: "setsum (\<lambda>v. u v *\<^sub>R v) S = a"
+ and SP: "S \<subseteq> P - {a}" and ua: "sum (\<lambda>v. u v *\<^sub>R v) S = a"
unfolding dependent_def span_explicit by blast
let ?S = "insert a S"
let ?u = "\<lambda>y. if y = a then - 1 else u y"
@@ -689,11 +689,11 @@
by blast
from fS SP aP have th0: "finite ?S" "?S \<subseteq> P" "?v \<in> ?S" "?u ?v \<noteq> 0"
by auto
- have s0: "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = 0"
+ have s0: "sum (\<lambda>v. ?u v *\<^sub>R v) ?S = 0"
using fS aS
apply simp
apply (subst (2) ua[symmetric])
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply auto
done
with th0 have ?rhs by fast
@@ -705,18 +705,18 @@
and SP: "S \<subseteq> P"
and vS: "v \<in> S"
and uv: "u v \<noteq> 0"
- and u: "setsum (\<lambda>v. u v *\<^sub>R v) S = 0"
+ and u: "sum (\<lambda>v. u v *\<^sub>R v) S = 0"
let ?a = v
let ?S = "S - {v}"
let ?u = "\<lambda>i. (- u i) / u v"
have th0: "?a \<in> P" "finite ?S" "?S \<subseteq> P"
using fS SP vS by auto
- have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S =
- setsum (\<lambda>v. (- (inverse (u ?a))) *\<^sub>R (u v *\<^sub>R v)) S - ?u v *\<^sub>R v"
- using fS vS uv by (simp add: setsum_diff1 field_simps)
+ have "sum (\<lambda>v. ?u v *\<^sub>R v) ?S =
+ sum (\<lambda>v. (- (inverse (u ?a))) *\<^sub>R (u v *\<^sub>R v)) S - ?u v *\<^sub>R v"
+ using fS vS uv by (simp add: sum_diff1 field_simps)
also have "\<dots> = ?a"
- unfolding scaleR_right.setsum [symmetric] u using uv by simp
- finally have "setsum (\<lambda>v. ?u v *\<^sub>R v) ?S = ?a" .
+ unfolding scaleR_right.sum [symmetric] u using uv by simp
+ finally have "sum (\<lambda>v. ?u v *\<^sub>R v) ?S = ?a" .
with th0 have ?lhs
unfolding dependent_def span_explicit
apply -
@@ -740,7 +740,7 @@
by (force simp: dependent_explicit)
with assms show ?rhs
apply (rule_tac x="\<lambda>v. if v \<in> T then u v else 0" in exI)
- apply (auto simp: setsum.mono_neutral_right)
+ apply (auto simp: sum.mono_neutral_right)
done
next
assume ?rhs with assms show ?lhs
@@ -753,7 +753,7 @@
apply safe
subgoal for x S u
by (intro exI[of _ "\<lambda>x. if x \<in> S then u x else 0"])
- (auto intro!: setsum.mono_neutral_cong_right)
+ (auto intro!: sum.mono_neutral_cong_right)
apply auto
done
@@ -764,8 +764,8 @@
apply safe
subgoal for S u v
apply (intro exI[of _ "\<lambda>x. if x \<in> S then u x else 0"])
- apply (subst setsum.mono_neutral_cong_left[where T=S])
- apply (auto intro!: setsum.mono_neutral_cong_right cong: rev_conj_cong)
+ apply (subst sum.mono_neutral_cong_left[where T=S])
+ apply (auto intro!: sum.mono_neutral_cong_right cong: rev_conj_cong)
done
apply auto
done
@@ -794,13 +794,13 @@
then show "finite {x. X x - Y x \<noteq> 0}" "{x. X x - Y x \<noteq> 0} \<subseteq> B"
using X Y by (auto dest: finite_subset)
then have "(\<Sum>x | X x - Y x \<noteq> 0. (X x - Y x) *\<^sub>R x) = (\<Sum>v\<in>{S. X S \<noteq> 0} \<union> {S. Y S \<noteq> 0}. (X v - Y v) *\<^sub>R v)"
- using X Y by (intro setsum.mono_neutral_cong_left) auto
+ using X Y by (intro sum.mono_neutral_cong_left) auto
also have "\<dots> = (\<Sum>v\<in>{S. X S \<noteq> 0} \<union> {S. Y S \<noteq> 0}. X v *\<^sub>R v) - (\<Sum>v\<in>{S. X S \<noteq> 0} \<union> {S. Y S \<noteq> 0}. Y v *\<^sub>R v)"
- by (simp add: scaleR_diff_left setsum_subtractf assms)
+ by (simp add: scaleR_diff_left sum_subtractf assms)
also have "(\<Sum>v\<in>{S. X S \<noteq> 0} \<union> {S. Y S \<noteq> 0}. X v *\<^sub>R v) = (\<Sum>v\<in>{S. X S \<noteq> 0}. X v *\<^sub>R v)"
- using X Y by (intro setsum.mono_neutral_cong_right) auto
+ using X Y by (intro sum.mono_neutral_cong_right) auto
also have "(\<Sum>v\<in>{S. X S \<noteq> 0} \<union> {S. Y S \<noteq> 0}. Y v *\<^sub>R v) = (\<Sum>v\<in>{S. Y S \<noteq> 0}. Y v *\<^sub>R v)"
- using X Y by (intro setsum.mono_neutral_cong_right) auto
+ using X Y by (intro sum.mono_neutral_cong_right) auto
finally show "(\<Sum>x | X x - Y x \<noteq> 0. (X x - Y x) *\<^sub>R x) = 0"
using assms by simp
qed
@@ -927,7 +927,7 @@
lemma span_finite:
assumes fS: "finite S"
- shows "span S = {y. \<exists>u. setsum (\<lambda>v. u v *\<^sub>R v) S = y}"
+ shows "span S = {y. \<exists>u. sum (\<lambda>v. u v *\<^sub>R v) S = y}"
(is "_ = ?rhs")
proof -
{
@@ -935,18 +935,18 @@
assume y: "y \<in> span S"
from y obtain S' u where fS': "finite S'"
and SS': "S' \<subseteq> S"
- and u: "setsum (\<lambda>v. u v *\<^sub>R v) S' = y"
+ and u: "sum (\<lambda>v. u v *\<^sub>R v) S' = y"
unfolding span_explicit by blast
let ?u = "\<lambda>x. if x \<in> S' then u x else 0"
- have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = setsum (\<lambda>v. u v *\<^sub>R v) S'"
- using SS' fS by (auto intro!: setsum.mono_neutral_cong_right)
- then have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = y" by (metis u)
+ have "sum (\<lambda>v. ?u v *\<^sub>R v) S = sum (\<lambda>v. u v *\<^sub>R v) S'"
+ using SS' fS by (auto intro!: sum.mono_neutral_cong_right)
+ then have "sum (\<lambda>v. ?u v *\<^sub>R v) S = y" by (metis u)
then have "y \<in> ?rhs" by auto
}
moreover
{
fix y u
- assume u: "setsum (\<lambda>v. u v *\<^sub>R v) S = y"
+ assume u: "sum (\<lambda>v. u v *\<^sub>R v) S = y"
then have "y \<in> span S" using fS unfolding span_explicit by auto
}
ultimately show ?thesis by blast
@@ -970,10 +970,10 @@
proof (rule independentD_unique)
have "(\<Sum>z | X x z + X y z \<noteq> 0. (X x z + X y z) *\<^sub>R z)
= (\<Sum>z\<in>{z. X x z \<noteq> 0} \<union> {z. X y z \<noteq> 0}. (X x z + X y z) *\<^sub>R z)"
- by (intro setsum.mono_neutral_cong_left) (auto intro: X)
+ by (intro sum.mono_neutral_cong_left) (auto intro: X)
also have "\<dots> = (\<Sum>z\<in>{z. X x z \<noteq> 0}. X x z *\<^sub>R z) + (\<Sum>z\<in>{z. X y z \<noteq> 0}. X y z *\<^sub>R z)"
- by (auto simp add: scaleR_add_left setsum.distrib
- intro!: arg_cong2[where f="op +"] setsum.mono_neutral_cong_right X)
+ by (auto simp add: scaleR_add_left sum.distrib
+ intro!: arg_cong2[where f="op +"] sum.mono_neutral_cong_right X)
also have "\<dots> = x + y"
by (simp add: X(3)[symmetric])
also have "\<dots> = (\<Sum>z | X (x + y) z \<noteq> 0. X (x + y) z *\<^sub>R z)"
@@ -994,7 +994,7 @@
"finite {z. c * X x z \<noteq> 0}" "{z. c * X x z \<noteq> 0} \<subseteq> B' "
using X(1,2) by auto
show "(\<Sum>z | X (c *\<^sub>R x) z \<noteq> 0. X (c *\<^sub>R x) z *\<^sub>R z) = (\<Sum>z | c * X x z \<noteq> 0. (c * X x z) *\<^sub>R z)"
- unfolding scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric]
+ unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric]
by (cases "c = 0") (auto simp: X(3)[symmetric])
qed
@@ -1013,14 +1013,14 @@
fix x y
have *: "(\<Sum>z | X x z + X y z \<noteq> 0. (X x z + X y z) *\<^sub>R f' z)
= (\<Sum>z\<in>{z. X x z \<noteq> 0} \<union> {z. X y z \<noteq> 0}. (X x z + X y z) *\<^sub>R f' z)"
- by (intro setsum.mono_neutral_cong_left) (auto intro: X)
+ by (intro sum.mono_neutral_cong_left) (auto intro: X)
show "g (x + y) = g x + g y"
unfolding g_def X_add *
- by (auto simp add: scaleR_add_left setsum.distrib
- intro!: arg_cong2[where f="op +"] setsum.mono_neutral_cong_right X)
+ by (auto simp add: scaleR_add_left sum.distrib
+ intro!: arg_cong2[where f="op +"] sum.mono_neutral_cong_right X)
next
show "g (r *\<^sub>R x) = r *\<^sub>R g x" for r x
- by (auto simp add: g_def X_cmult scaleR_setsum_right intro!: setsum.mono_neutral_cong_left X)
+ by (auto simp add: g_def X_cmult scaleR_sum_right intro!: sum.mono_neutral_cong_left X)
qed
moreover have "\<forall>x\<in>B. g x = f x"
using \<open>B \<subseteq> B'\<close> by (auto simp: g_f' f'_def)
@@ -1436,9 +1436,9 @@
shows "(f x) \<bullet> j = (\<Sum>i\<in>Basis. (x\<bullet>i) * (f i\<bullet>j))" (is "?lhs = ?rhs")
proof -
have "?rhs = (\<Sum>i\<in>Basis. (x\<bullet>i) *\<^sub>R (f i))\<bullet>j"
- by (simp add: inner_setsum_left)
+ by (simp add: inner_sum_left)
then show ?thesis
- unfolding linear_setsum_mul[OF lf, symmetric]
+ unfolding linear_sum_mul[OF lf, symmetric]
unfolding euclidean_representation ..
qed
@@ -1475,30 +1475,30 @@
lemma norm_triangle_lt: "norm x + norm y < e \<Longrightarrow> norm (x + y) < e"
by (rule norm_triangle_ineq [THEN le_less_trans])
-lemma setsum_clauses:
- shows "setsum f {} = 0"
- and "finite S \<Longrightarrow> setsum f (insert x S) = (if x \<in> S then setsum f S else f x + setsum f S)"
+lemma sum_clauses:
+ shows "sum f {} = 0"
+ and "finite S \<Longrightarrow> sum f (insert x S) = (if x \<in> S then sum f S else f x + sum f S)"
by (auto simp add: insert_absorb)
-lemma setsum_norm_le:
+lemma sum_norm_le:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
- shows "norm (setsum f S) \<le> setsum g S"
- by (rule order_trans [OF norm_setsum setsum_mono]) (simp add: fg)
-
-lemma setsum_norm_bound:
+ shows "norm (sum f S) \<le> sum g S"
+ by (rule order_trans [OF norm_sum sum_mono]) (simp add: fg)
+
+lemma sum_norm_bound:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes K: "\<forall>x \<in> S. norm (f x) \<le> K"
- shows "norm (setsum f S) \<le> of_nat (card S) * K"
- using setsum_norm_le[OF K] setsum_constant[symmetric]
+ shows "norm (sum f S) \<le> of_nat (card S) * K"
+ using sum_norm_le[OF K] sum_constant[symmetric]
by simp
-lemma setsum_group:
+lemma sum_group:
assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
- shows "setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) T = setsum g S"
- apply (subst setsum_image_gen[OF fS, of g f])
- apply (rule setsum.mono_neutral_right[OF fT fST])
- apply (auto intro: setsum.neutral)
+ shows "sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) T = sum g S"
+ apply (subst sum_image_gen[OF fS, of g f])
+ apply (rule sum.mono_neutral_right[OF fT fST])
+ apply (auto intro: sum.neutral)
done
lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"
@@ -1557,11 +1557,11 @@
by (auto simp: pairwise_def orthogonal_clauses)
lemma orthogonal_rvsum:
- "\<lbrakk>finite s; \<And>y. y \<in> s \<Longrightarrow> orthogonal x (f y)\<rbrakk> \<Longrightarrow> orthogonal x (setsum f s)"
+ "\<lbrakk>finite s; \<And>y. y \<in> s \<Longrightarrow> orthogonal x (f y)\<rbrakk> \<Longrightarrow> orthogonal x (sum f s)"
by (induction s rule: finite_induct) (auto simp: orthogonal_clauses)
lemma orthogonal_lvsum:
- "\<lbrakk>finite s; \<And>x. x \<in> s \<Longrightarrow> orthogonal (f x) y\<rbrakk> \<Longrightarrow> orthogonal (setsum f s) y"
+ "\<lbrakk>finite s; \<And>x. x \<in> s \<Longrightarrow> orthogonal (f x) y\<rbrakk> \<Longrightarrow> orthogonal (sum f s) y"
by (induction s rule: finite_induct) (auto simp: orthogonal_clauses)
lemma norm_add_Pythagorean:
@@ -1574,15 +1574,15 @@
by (simp add: power2_norm_eq_inner)
qed
-lemma norm_setsum_Pythagorean:
+lemma norm_sum_Pythagorean:
assumes "finite I" "pairwise (\<lambda>i j. orthogonal (f i) (f j)) I"
- shows "(norm (setsum f I))\<^sup>2 = (\<Sum>i\<in>I. (norm (f i))\<^sup>2)"
+ shows "(norm (sum f I))\<^sup>2 = (\<Sum>i\<in>I. (norm (f i))\<^sup>2)"
using assms
proof (induction I rule: finite_induct)
case empty then show ?case by simp
next
case (insert x I)
- then have "orthogonal (f x) (setsum f I)"
+ then have "orthogonal (f x) (sum f I)"
by (metis pairwise_insert orthogonal_rvsum)
with insert show ?case
by (simp add: pairwise_insert norm_add_Pythagorean)
@@ -1630,25 +1630,25 @@
lemma bilinear_rsub: "bilinear h \<Longrightarrow> h z (x - y) = h z x - h z y"
using bilinear_radd [of h _ x "- y"] by (simp add: bilinear_rneg)
-lemma bilinear_setsum:
+lemma bilinear_sum:
assumes bh: "bilinear h"
and fS: "finite S"
and fT: "finite T"
- shows "h (setsum f S) (setsum g T) = setsum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "
+ shows "h (sum f S) (sum g T) = sum (\<lambda>(i,j). h (f i) (g j)) (S \<times> T) "
proof -
- have "h (setsum f S) (setsum g T) = setsum (\<lambda>x. h (f x) (setsum g T)) S"
- apply (rule linear_setsum[unfolded o_def])
+ have "h (sum f S) (sum g T) = sum (\<lambda>x. h (f x) (sum g T)) S"
+ apply (rule linear_sum[unfolded o_def])
using bh fS
apply (auto simp add: bilinear_def)
done
- also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"
- apply (rule setsum.cong, simp)
- apply (rule linear_setsum[unfolded o_def])
+ also have "\<dots> = sum (\<lambda>x. sum (\<lambda>y. h (f x) (g y)) T) S"
+ apply (rule sum.cong, simp)
+ apply (rule linear_sum[unfolded o_def])
using bh fT
apply (auto simp add: bilinear_def)
done
finally show ?thesis
- unfolding setsum.cartesian_product .
+ unfolding sum.cartesian_product .
qed
@@ -1694,10 +1694,10 @@
have "f x \<bullet> y = f (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i) \<bullet> y"
by (simp add: euclidean_representation)
also have "\<dots> = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R f i) \<bullet> y"
- unfolding linear_setsum[OF lf]
+ unfolding linear_sum[OF lf]
by (simp add: linear_cmul[OF lf])
finally show "f x \<bullet> y = x \<bullet> ?w"
- by (simp add: inner_setsum_left inner_setsum_right mult.commute)
+ by (simp add: inner_sum_left inner_sum_right mult.commute)
qed
then show ?thesis
unfolding adjoint_def choice_iff
@@ -1876,7 +1876,7 @@
apply simp
apply clarify
apply (drule_tac f="inner a" in arg_cong)
- apply (simp add: inner_Basis inner_setsum_right eq_commute)
+ apply (simp add: inner_Basis inner_sum_right eq_commute)
done
lemma span_Basis [simp]: "span Basis = UNIV"
@@ -1897,27 +1897,27 @@
lemma norm_le_l1: "norm x \<le> (\<Sum>b\<in>Basis. \<bar>x \<bullet> b\<bar>)"
apply (subst euclidean_representation[of x, symmetric])
- apply (rule order_trans[OF norm_setsum])
- apply (auto intro!: setsum_mono)
+ apply (rule order_trans[OF norm_sum])
+ apply (auto intro!: sum_mono)
done
-lemma setsum_norm_allsubsets_bound:
+lemma sum_norm_allsubsets_bound:
fixes f :: "'a \<Rightarrow> 'n::euclidean_space"
assumes fP: "finite P"
- and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (setsum f Q) \<le> e"
+ and fPs: "\<And>Q. Q \<subseteq> P \<Longrightarrow> norm (sum f Q) \<le> e"
shows "(\<Sum>x\<in>P. norm (f x)) \<le> 2 * real DIM('n) * e"
proof -
have "(\<Sum>x\<in>P. norm (f x)) \<le> (\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>)"
- by (rule setsum_mono) (rule norm_le_l1)
+ by (rule sum_mono) (rule norm_le_l1)
also have "(\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>) = (\<Sum>b\<in>Basis. \<Sum>x\<in>P. \<bar>f x \<bullet> b\<bar>)"
- by (rule setsum.commute)
+ by (rule sum.commute)
also have "\<dots> \<le> of_nat (card (Basis :: 'n set)) * (2 * e)"
- proof (rule setsum_bounded_above)
+ proof (rule sum_bounded_above)
fix i :: 'n
assume i: "i \<in> Basis"
have "norm (\<Sum>x\<in>P. \<bar>f x \<bullet> i\<bar>) \<le>
norm ((\<Sum>x\<in>P \<inter> - {x. f x \<bullet> i < 0}. f x) \<bullet> i) + norm ((\<Sum>x\<in>P \<inter> {x. f x \<bullet> i < 0}. f x) \<bullet> i)"
- by (simp add: abs_real_def setsum.If_cases[OF fP] setsum_negf norm_triangle_ineq4 inner_setsum_left
+ by (simp add: abs_real_def sum.If_cases[OF fP] sum_negf norm_triangle_ineq4 inner_sum_left
del: real_norm_def)
also have "\<dots> \<le> e + e"
unfolding real_norm_def
@@ -1943,9 +1943,9 @@
let ?g = "\<lambda>b. (x \<bullet> b) *\<^sub>R f b"
have "norm (f x) = norm (f (\<Sum>b\<in>Basis. (x \<bullet> b) *\<^sub>R b))"
unfolding euclidean_representation ..
- also have "\<dots> = norm (setsum ?g Basis)"
- by (simp add: linear_setsum [OF lf] linear_cmul [OF lf])
- finally have th0: "norm (f x) = norm (setsum ?g Basis)" .
+ also have "\<dots> = norm (sum ?g Basis)"
+ by (simp add: linear_sum [OF lf] linear_cmul [OF lf])
+ finally have th0: "norm (f x) = norm (sum ?g Basis)" .
have th: "\<forall>b\<in>Basis. norm (?g b) \<le> norm (f b) * norm x"
proof
fix i :: 'a
@@ -1958,9 +1958,9 @@
apply (auto simp add: field_simps)
done
qed
- from setsum_norm_le[of _ ?g, OF th]
+ from sum_norm_le[of _ ?g, OF th]
show "norm (f x) \<le> ?B * norm x"
- unfolding th0 setsum_distrib_right by metis
+ unfolding th0 sum_distrib_right by metis
qed
qed
@@ -2012,17 +2012,17 @@
proof (clarify intro!: exI[of _ "\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)"])
fix x :: 'm
fix y :: 'n
- have "norm (h x y) = norm (h (setsum (\<lambda>i. (x \<bullet> i) *\<^sub>R i) Basis) (setsum (\<lambda>i. (y \<bullet> i) *\<^sub>R i) Basis))"
+ have "norm (h x y) = norm (h (sum (\<lambda>i. (x \<bullet> i) *\<^sub>R i) Basis) (sum (\<lambda>i. (y \<bullet> i) *\<^sub>R i) Basis))"
apply (subst euclidean_representation[where 'a='m])
apply (subst euclidean_representation[where 'a='n])
apply rule
done
- also have "\<dots> = norm (setsum (\<lambda> (i,j). h ((x \<bullet> i) *\<^sub>R i) ((y \<bullet> j) *\<^sub>R j)) (Basis \<times> Basis))"
- unfolding bilinear_setsum[OF bh finite_Basis finite_Basis] ..
+ also have "\<dots> = norm (sum (\<lambda> (i,j). h ((x \<bullet> i) *\<^sub>R i) ((y \<bullet> j) *\<^sub>R j)) (Basis \<times> Basis))"
+ unfolding bilinear_sum[OF bh finite_Basis finite_Basis] ..
finally have th: "norm (h x y) = \<dots>" .
show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
- apply (auto simp add: setsum_distrib_right th setsum.cartesian_product)
- apply (rule setsum_norm_le)
+ apply (auto simp add: sum_distrib_right th sum.cartesian_product)
+ apply (rule sum_norm_le)
apply simp
apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
field_simps simp del: scaleR_scaleR)
@@ -2353,7 +2353,7 @@
from \<open>\<exists>C. finite C \<and> card C \<le> card B \<and> span C = span B \<and> pairwise orthogonal C\<close>
obtain C where C: "finite C" "card C \<le> card B"
"span C = span B" "pairwise orthogonal C" by blast
- let ?a = "a - setsum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C"
+ let ?a = "a - sum (\<lambda>x. (x \<bullet> a / (x \<bullet> x)) *\<^sub>R x) C"
let ?C = "insert ?a C"
from C(1) have fC: "finite ?C"
by simp
@@ -2367,7 +2367,7 @@
apply (simp only: scaleR_right_diff_distrib th0)
apply (rule span_add_eq)
apply (rule span_mul)
- apply (rule span_setsum)
+ apply (rule span_sum)
apply (rule span_mul)
apply (rule span_superset)
apply assumption
@@ -2384,10 +2384,10 @@
using C by simp
have "orthogonal ?a y"
unfolding orthogonal_def
- unfolding inner_diff inner_setsum_left right_minus_eq
- unfolding setsum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>]
+ unfolding inner_diff inner_sum_left right_minus_eq
+ unfolding sum.remove [OF \<open>finite C\<close> \<open>y \<in> C\<close>]
apply (clarsimp simp add: inner_commute[of y a])
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply clarsimp
apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
using \<open>y \<in> C\<close> by auto
@@ -2446,10 +2446,10 @@
from span_mono[OF B(2)] span_mono[OF B(3)]
have sSB: "span S = span B"
by (simp add: span_span)
- let ?a = "a - setsum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B"
- have "setsum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S"
+ let ?a = "a - sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B"
+ have "sum (\<lambda>b. (a \<bullet> b / (b \<bullet> b)) *\<^sub>R b) B \<in> span S"
unfolding sSB
- apply (rule span_setsum)
+ apply (rule span_sum)
apply (rule span_mul)
apply (rule span_superset)
apply assumption
@@ -2471,10 +2471,10 @@
have "?a \<bullet> x = 0"
apply (subst B')
using fB fth
- unfolding setsum_clauses(2)[OF fth]
+ unfolding sum_clauses(2)[OF fth]
apply simp unfolding inner_simps
- apply (clarsimp simp add: inner_add inner_setsum_left)
- apply (rule setsum.neutral, rule ballI)
+ apply (clarsimp simp add: inner_add inner_sum_left)
+ apply (rule sum.neutral, rule ballI)
apply (simp only: inner_commute)
apply (auto simp add: x field_simps
intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])
@@ -3016,7 +3016,7 @@
lemma (in euclidean_space) euclidean_inner: "inner x y = (\<Sum>b\<in>Basis. (x \<bullet> b) * (y \<bullet> b))"
by (subst (1 2) euclidean_representation [symmetric])
- (simp add: inner_setsum_right inner_Basis ac_simps)
+ (simp add: inner_sum_right inner_Basis ac_simps)
lemma norm_le_infnorm:
fixes x :: "'a::euclidean_space"
@@ -3033,7 +3033,7 @@
unfolding power_mult_distrib d2
apply (subst euclidean_inner)
apply (subst power2_abs[symmetric])
- apply (rule order_trans[OF setsum_bounded_above[where K="\<bar>infnorm x\<bar>\<^sup>2"]])
+ apply (rule order_trans[OF sum_bounded_above[where K="\<bar>infnorm x\<bar>\<^sup>2"]])
apply (auto simp add: power2_eq_square[symmetric])
apply (subst power2_abs[symmetric])
apply (rule power_mono)
--- a/src/HOL/Analysis/Measure_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Measure_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -21,7 +21,7 @@
have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
- by (auto simp: setsum.If_cases)
+ by (auto simp: sum.If_cases)
moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
proof (rule SUP_eqI)
fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
@@ -41,7 +41,7 @@
show ?thesis using * by simp
qed simp
-lemma setsum_indicator_disjoint_family:
+lemma sum_indicator_disjoint_family:
fixes f :: "'d \<Rightarrow> 'e::semiring_1"
assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
@@ -51,7 +51,7 @@
by auto
thus ?thesis
unfolding indicator_def
- by (simp add: if_distrib setsum.If_cases[OF \<open>finite P\<close>])
+ by (simp add: if_distrib sum.If_cases[OF \<open>finite P\<close>])
qed
text \<open>
@@ -302,7 +302,7 @@
from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
by (rule suminf_finite) auto
also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
- using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
+ using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto
also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto
also have "\<dots> = \<mu> (\<Union>i. F i)"
@@ -476,7 +476,7 @@
"A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)"
using plus_emeasure[of A M "B - A"] by auto
-lemma setsum_emeasure:
+lemma sum_emeasure:
"F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
(\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
by (metis sets.additive_sum emeasure_positive emeasure_additive)
@@ -704,20 +704,20 @@
"A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
by (rule emeasure_insert)
-lemma emeasure_eq_setsum_singleton:
+lemma emeasure_eq_sum_singleton:
assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
- using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
+ using sum_emeasure[of "\<lambda>x. {x}" S M] assms
by (auto simp: disjoint_family_on_def subset_eq)
-lemma setsum_emeasure_cover:
+lemma sum_emeasure_cover:
assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
assumes disj: "disjoint_family_on B S"
shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
proof -
have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
- proof (rule setsum_emeasure)
+ proof (rule sum_emeasure)
show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
using \<open>disjoint_family_on B S\<close>
unfolding disjoint_family_on_def by auto
@@ -749,11 +749,11 @@
fix X assume "X \<in> sets M"
then have X: "X \<subseteq> A" by auto
then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
- using \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+ using \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset)
also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
- using X eq by (auto intro!: setsum.cong)
+ using X eq by (auto intro!: sum.cong)
also have "\<dots> = emeasure N X"
- using X \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+ using X \<open>finite A\<close> by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset)
finally show "emeasure M X = emeasure N X" .
qed simp
@@ -1320,7 +1320,7 @@
have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
using F by (auto intro!: emeasure_subadditive_finite)
also have "\<dots> < \<infinity>"
- using F by (auto simp: setsum_Pinfty less_top)
+ using F by (auto simp: sum_Pinfty less_top)
finally show ?thesis by simp
qed
show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
@@ -1517,7 +1517,7 @@
"finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow>
measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
by (induction S rule: finite_induct)
- (auto simp: disjoint_family_on_insert measure_Union setsum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
+ (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
lemma measure_Diff:
assumes finite: "emeasure M A \<noteq> \<infinity>"
@@ -1576,7 +1576,7 @@
show ?thesis
using emeasure_subadditive_finite[OF A] fin
unfolding emeasure_eq_ennreal_measure[OF *]
- by (simp_all add: setsum_nonneg emeasure_eq_ennreal_measure)
+ by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure)
qed
lemma measure_subadditive_countably:
@@ -1611,11 +1611,11 @@
lemma measure_Diff_null_set: "A \<in> sets M \<Longrightarrow> B \<in> null_sets M \<Longrightarrow> measure M (A - B) = measure M A"
by (simp add: measure_def emeasure_Diff_null_set)
-lemma measure_eq_setsum_singleton:
+lemma measure_eq_sum_singleton:
"finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
measure M S = (\<Sum>x\<in>S. measure M {x})"
- using emeasure_eq_setsum_singleton[of S M]
- by (intro measure_eq_emeasure_eq_ennreal) (auto simp: setsum_nonneg emeasure_eq_ennreal_measure)
+ using emeasure_eq_sum_singleton[of S M]
+ by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure)
lemma Lim_measure_incseq:
assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
@@ -1886,10 +1886,10 @@
by (rule measure_subadditive_countably)
(simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)
-lemma (in finite_measure) finite_measure_eq_setsum_singleton:
+lemma (in finite_measure) finite_measure_eq_sum_singleton:
assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
- using measure_eq_setsum_singleton[OF assms] by simp
+ using measure_eq_sum_singleton[OF assms] by simp
lemma (in finite_measure) finite_Lim_measure_incseq:
assumes A: "range A \<subseteq> sets M" "incseq A"
@@ -1985,10 +1985,10 @@
from someI_ex[OF this] assms
have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
have "measure M s = (\<Sum> x \<in> s. measure M {x})"
- using finite_measure_eq_setsum_singleton[OF s] by simp
+ using finite_measure_eq_sum_singleton[OF s] by simp
also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
- using setsum_constant assms by simp
+ using sum_constant assms by simp
finally show ?thesis by simp
qed simp
@@ -3581,7 +3581,7 @@
using \<open>?M \<noteq> 0\<close>
by (simp add: \<open>card X = Suc (Suc n)\<close> field_simps less_le)
also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
- by (rule setsum_mono) fact
+ by (rule sum_mono) fact
also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
using singleton_sets \<open>finite X\<close>
by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
--- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Mon Oct 17 14:37:32 2016 +0200
@@ -68,9 +68,9 @@
proof -
have "?r = (\<Sum>y \<in> f ` space M.
(if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
- by (auto intro!: setsum.cong)
+ by (auto intro!: sum.cong)
also have "... = f x * indicator (f -` {f x} \<inter> space M) x"
- using assms by (auto dest: simple_functionD simp: setsum.delta)
+ using assms by (auto dest: simple_functionD simp: sum.delta)
also have "... = f x" using x by (auto simp: indicator_def)
finally show ?thesis by auto
qed
@@ -188,7 +188,7 @@
and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
-lemma simple_function_setsum[intro, simp]:
+lemma simple_function_sum[intro, simp]:
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
proof cases
@@ -362,10 +362,10 @@
case (insert x S)
{ fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
x * indicator (u -` {x} \<inter> space M) z = 0"
- using insert by (subst setsum_eq_0_iff) (auto simp: indicator_def) }
+ using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) }
note disj = this
from insert show ?case
- by (auto intro!: add mult set simple_functionD u simple_function_setsum disj)
+ by (auto intro!: add mult set simple_functionD u simple_function_sum disj)
qed
qed fact
qed
@@ -530,7 +530,7 @@
(\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M.
if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
unfolding simple_integral_def
- proof (safe intro!: setsum.cong ennreal_mult_left_cong)
+ proof (safe intro!: sum.cong ennreal_mult_left_cong)
fix y assume y: "y \<in> space M" "f y \<noteq> 0"
have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
@@ -543,17 +543,17 @@
by (rule rev_finite_subset) auto
then show "emeasure M (f -` {f y} \<inter> space M) =
(\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
- apply (simp add: setsum.If_cases)
- apply (subst setsum_emeasure)
+ apply (simp add: sum.If_cases)
+ apply (subst sum_emeasure)
apply (auto simp: disjoint_family_on_def eq)
done
qed
also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
- by (auto intro!: setsum.cong simp: setsum_distrib_left)
+ by (auto intro!: sum.cong simp: sum_distrib_left)
also have "\<dots> = ?r"
- by (subst setsum.commute)
- (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
+ by (subst sum.commute)
+ (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq)
finally show "integral\<^sup>S M f = ?r" .
qed
@@ -566,7 +566,7 @@
by (intro simple_function_partition) (auto intro: f g)
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
- using assms(2,4) by (auto intro!: setsum.cong distrib_right simp: setsum.distrib[symmetric])
+ using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric])
also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
by (intro simple_function_partition[symmetric]) (auto intro: f g)
also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
@@ -574,14 +574,14 @@
finally show ?thesis .
qed
-lemma simple_integral_setsum[simp]:
+lemma simple_integral_sum[simp]:
assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))"
proof cases
assume "finite P"
from this assms show ?thesis
- by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
+ by induct (auto simp: simple_function_sum simple_integral_add sum_nonneg)
qed auto
lemma simple_integral_mult[simp]:
@@ -592,7 +592,7 @@
using f by (intro simple_function_partition) auto
also have "\<dots> = c * integral\<^sup>S M f"
using f unfolding simple_integral_def
- by (subst setsum_distrib_left) (auto simp: mult.assoc Int_def conj_commute)
+ by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute)
finally show ?thesis .
qed
@@ -605,7 +605,7 @@
have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
using f g by (intro simple_function_partition) auto
also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
- proof (clarsimp intro!: setsum_mono)
+ proof (clarsimp intro!: sum_mono)
fix x assume "x \<in> space M"
let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)"
show "f x * ?M \<le> g x * ?M"
@@ -664,14 +664,14 @@
using assms by (intro simple_function_partition) auto
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M.
if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
- by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong)
+ by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] sum.cong)
also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
- using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
+ using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
- by (subst setsum.reindex [of fst]) (auto simp: inj_on_def)
+ by (subst sum.reindex [of fst]) (auto simp: inj_on_def)
also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
using A[THEN sets.sets_into_space]
- by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
+ by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
finally show ?thesis .
qed
@@ -848,7 +848,7 @@
let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))"
unfolding simple_integral_def
- proof (intro setsum.cong ennreal_mult_left_cong refl)
+ proof (intro sum.cong ennreal_mult_left_cong refl)
fix c assume "c \<in> ?au ` space M" "c \<noteq> 0"
{ fix x' assume x': "x' \<in> space M" "?au x' = c"
with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'"
@@ -866,7 +866,7 @@
qed
also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))"
unfolding SUP_mult_left_ennreal using f
- by (intro ennreal_SUP_setsum[symmetric])
+ by (intro ennreal_SUP_sum[symmetric])
(auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))"
proof (intro SUP_subset_mono order_refl)
@@ -874,7 +874,7 @@
have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) =
(\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
by (subst simple_integral_indicator)
- (auto intro!: setsum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
+ (auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
by (rule nn_integral_eq_simple_integral[symmetric]) simp
also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)"
@@ -1031,7 +1031,7 @@
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
using nn_integral_linear[of f M g 1] by simp
-lemma nn_integral_setsum:
+lemma nn_integral_sum:
"(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)
@@ -1044,10 +1044,10 @@
have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
by (rule suminf_eq_SUP)
also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
- unfolding nn_integral_setsum[OF f] ..
+ unfolding nn_integral_sum[OF f] ..
also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
- (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
+ (elim AE_mp, auto simp: sum_nonneg simp del: sum_lessThan_Suc intro!: AE_I2 sum_mono3)
also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP)
finally show ?thesis by simp
@@ -1534,7 +1534,7 @@
by (simp add: simple_function_def)
show "integral\<^sup>S M x \<le> integral\<^sup>S N x" for x
using le_measureD3[OF \<open>M \<le> N\<close>]
- by (auto simp add: simple_integral_def intro!: setsum_mono mult_mono)
+ by (auto simp add: simple_integral_def intro!: sum_mono mult_mono)
qed
lemma nn_integral_empty:
@@ -1587,17 +1587,17 @@
have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
(\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
by (auto intro!: nn_integral_cong
- simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less)
+ simp add: indicator_def if_distrib sum.If_cases[OF A] max_def le_less)
also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
- by (subst nn_integral_setsum) (simp_all add: AE_count_space less_imp_le)
+ by (subst nn_integral_sum) (simp_all add: AE_count_space less_imp_le)
also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
- by (auto intro!: setsum.cong simp: one_ennreal_def[symmetric] max_def)
+ by (auto intro!: sum.cong simp: one_ennreal_def[symmetric] max_def)
finally show ?thesis by (simp add: max.absorb2)
qed
lemma nn_integral_count_space_finite:
"finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
- by (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
+ by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le)
lemma nn_integral_count_space':
assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B"
@@ -1607,7 +1607,7 @@
using assms(2,3)
by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le)
also have "\<dots> = (\<Sum>a\<in>A. f a)"
- using assms by (intro setsum.mono_neutral_cong_left) (auto simp: less_le)
+ using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le)
finally show ?thesis .
qed
@@ -1624,9 +1624,9 @@
shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
proof -
from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
- by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] setsum.If_cases)
+ by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] sum.If_cases)
also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
- by (subst nn_integral_setsum) auto
+ by (subst nn_integral_sum) auto
finally show ?thesis .
qed
@@ -1680,7 +1680,7 @@
shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)"
proof cases
assume "finite I" then show ?thesis
- by (simp add: nn_integral_count_space_finite nn_integral_setsum)
+ by (simp add: nn_integral_count_space_finite nn_integral_sum)
next
assume "infinite I"
then have [simp]: "I \<noteq> {}"
@@ -1823,7 +1823,7 @@
then obtain m' where "m' < m" and m': "(SUP i:?Y. f i) = f (I m')" by auto
have "I m' \<in> Y" using I by blast
have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)"
- proof(rule setsum_mono)
+ proof(rule sum_mono)
fix x
assume "x \<in> {..<m}"
hence "x < m" by simp
@@ -1920,7 +1920,7 @@
using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
split: split_indicator split_indicator_asm
- intro!: setsum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
+ intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
lemma nn_integral_restrict_space:
assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
@@ -2234,12 +2234,12 @@
lemma emeasure_point_measure_finite:
"finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
- by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
+ by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)
lemma emeasure_point_measure_finite2:
"X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
by (subst emeasure_point_measure)
- (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
+ (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le)
lemma null_sets_point_measure_iff:
"X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)"
@@ -2259,7 +2259,7 @@
lemma nn_integral_point_measure_finite:
"finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
- by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
+ by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le)
subsubsection \<open>Uniform measure\<close>
--- a/src/HOL/Analysis/Ordered_Euclidean_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Ordered_Euclidean_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -46,7 +46,7 @@
lemma inner_Basis_inf_left: "i \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"
- by (simp_all add: eucl_inf eucl_sup inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.delta
+ by (simp_all add: eucl_inf eucl_sup inner_sum_left inner_Basis if_distrib comm_monoid_add_class.sum.delta
cong: if_cong)
lemma inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x:X. f x) \<bullet> i = (INF x:X. f x \<bullet> i)"
@@ -65,7 +65,7 @@
shows "a\<bullet>i \<le> x\<bullet>i" "x\<bullet>i \<le> b\<bullet>i"
using assms
unfolding euclidean_inner[of a i] euclidean_inner[of x i] euclidean_inner[of b i]
- by (auto intro!: ordered_comm_monoid_add_class.setsum_mono mult_right_mono simp: eucl_le)
+ by (auto intro!: ordered_comm_monoid_add_class.sum_mono mult_right_mono simp: eucl_le)
lemma inner_nonneg_nonneg:
shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a \<bullet> b"
@@ -86,7 +86,7 @@
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
shows "Sup s = X"
using assms
- unfolding eucl_Sup euclidean_representation_setsum
+ unfolding eucl_Sup euclidean_representation_sum
by (auto intro!: conditionally_complete_lattice_class.cSup_eq_maximum)
lemma Inf_eq_minimum_componentwise:
@@ -95,7 +95,7 @@
assumes i_s: "\<And>b. b \<in> Basis \<Longrightarrow> (i b \<bullet> b) \<in> (\<lambda>x. x \<bullet> b) ` s"
shows "Inf s = X"
using assms
- unfolding eucl_Inf euclidean_representation_setsum
+ unfolding eucl_Inf euclidean_representation_sum
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
end
@@ -117,7 +117,7 @@
hence "Inf ?proj = x \<bullet> b"
by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum)
hence "x \<bullet> b = Inf X \<bullet> b"
- by (auto simp: eucl_Inf inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
+ by (auto simp: eucl_Inf inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Inf X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> x \<bullet> b \<le> y \<bullet> b)" by blast
qed
@@ -139,7 +139,7 @@
hence "Sup ?proj = x \<bullet> b"
by (auto intro!: cSup_eq_maximum)
hence "x \<bullet> b = Sup X \<bullet> b"
- by (auto simp: eucl_Sup[where 'a='a] inner_setsum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> setsum.delta
+ by (auto simp: eucl_Sup[where 'a='a] inner_sum_left inner_Basis if_distrib \<open>b \<in> Basis\<close> sum.delta
cong: if_cong)
with x show "\<exists>x. x \<in> X \<and> x \<bullet> b = Sup X \<bullet> b \<and> (\<forall>y. y \<in> X \<longrightarrow> y \<bullet> b \<le> x \<bullet> b)" by blast
qed
@@ -167,7 +167,7 @@
instance prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
by standard
- (auto intro!: add_mono simp add: euclidean_representation_setsum' Ball_def inner_prod_def
+ (auto intro!: add_mono simp add: euclidean_representation_sum' Ball_def inner_prod_def
in_Basis_prod_iff inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left Inf_prod_def
inner_Basis_SUP_left Sup_prod_def less_prod_def less_eq_prod_def eucl_le[where 'a='a]
eucl_le[where 'a='b] abs_prod_def abs_inner)
@@ -241,7 +241,7 @@
eucl_less (infix "<e" 50)
lemma One_nonneg: "0 \<le> (\<Sum>Basis::'a::ordered_euclidean_space)"
- by (auto intro: setsum_nonneg)
+ by (auto intro: sum_nonneg)
lemma
fixes a b::"'a::ordered_euclidean_space"
@@ -264,7 +264,7 @@
instance
apply standard
- unfolding euclidean_representation_setsum'
+ unfolding euclidean_representation_sum'
apply (auto simp: less_eq_vec_def inf_vec_def sup_vec_def Inf_vec_def Sup_vec_def inner_axis
Basis_vec_def inner_Basis_inf_left inner_Basis_sup_left inner_Basis_INF_left
inner_Basis_SUP_left eucl_le[where 'a='a] less_le_not_le abs_vec_def abs_inner)
--- a/src/HOL/Analysis/Path_Connected.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Path_Connected.thy Mon Oct 17 14:37:32 2016 +0200
@@ -5760,18 +5760,18 @@
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)"
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce
have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))"
- using linear_setsum [OF \<open>linear f\<close>] x by auto
+ using linear_sum [OF \<open>linear f\<close>] x by auto
also have "... = (\<Sum>v \<in> B. a v *\<^sub>R f v)"
- using \<open>linear f\<close> by (simp add: linear_setsum linear.scaleR)
+ using \<open>linear f\<close> by (simp add: linear_sum linear.scaleR)
also have "... = (\<Sum>v \<in> B. a v *\<^sub>R fb v)"
- by (simp add: ffb cong: setsum.cong)
+ by (simp add: ffb cong: sum.cong)
finally have "norm (f x)^2 = norm (\<Sum>v\<in>B. a v *\<^sub>R fb v)^2" by simp
also have "... = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)"
- apply (rule norm_setsum_Pythagorean [OF \<open>finite B\<close>])
+ apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>])
apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb])
done
also have "... = norm x ^2"
- by (simp add: x pairwise_ortho_scaleR Borth norm_setsum_Pythagorean [OF \<open>finite B\<close>])
+ by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>])
finally show ?thesis
by (simp add: norm_eq_sqrt_inner)
qed
@@ -5823,28 +5823,28 @@
obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)"
using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce
have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))"
- using linear_setsum [OF \<open>linear f\<close>] x by auto
+ using linear_sum [OF \<open>linear f\<close>] x by auto
also have "... = (\<Sum>v \<in> B. a v *\<^sub>R f v)"
- using \<open>linear f\<close> by (simp add: linear_setsum linear.scaleR)
+ using \<open>linear f\<close> by (simp add: linear_sum linear.scaleR)
also have "... = (\<Sum>v \<in> B. a v *\<^sub>R fb v)"
- by (simp add: ffb cong: setsum.cong)
+ by (simp add: ffb cong: sum.cong)
finally have *: "f x = (\<Sum>v\<in>B. a v *\<^sub>R fb v)" .
then have "(norm (f x))\<^sup>2 = (norm (\<Sum>v\<in>B. a v *\<^sub>R fb v))\<^sup>2" by simp
also have "... = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)"
- apply (rule norm_setsum_Pythagorean [OF \<open>finite B\<close>])
+ apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>])
apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb])
done
also have "... = (norm x)\<^sup>2"
- by (simp add: x pairwise_ortho_scaleR Borth norm_setsum_Pythagorean [OF \<open>finite B\<close>])
+ by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>])
finally show "norm (f x) = norm x"
by (simp add: norm_eq_sqrt_inner)
have "g (f x) = g (\<Sum>v\<in>B. a v *\<^sub>R fb v)" by (simp add: *)
also have "... = (\<Sum>v\<in>B. g (a v *\<^sub>R fb v))"
- using \<open>linear g\<close> by (simp add: linear_setsum linear.scaleR)
+ using \<open>linear g\<close> by (simp add: linear_sum linear.scaleR)
also have "... = (\<Sum>v\<in>B. a v *\<^sub>R g (fb v))"
by (simp add: \<open>linear g\<close> linear.scaleR)
also have "... = (\<Sum>v\<in>B. a v *\<^sub>R v)"
- apply (rule setsum.cong [OF refl])
+ apply (rule sum.cong [OF refl])
using \<open>bij_betw fb B C\<close> gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce
also have "... = x"
using x by blast
@@ -5857,14 +5857,14 @@
obtain a where x: "x = (\<Sum>v \<in> C. a v *\<^sub>R v)"
using \<open>finite C\<close> \<open>span C = T\<close> \<open>x \<in> T\<close> span_finite by fastforce
have "g x = (\<Sum>v \<in> C. g (a v *\<^sub>R v))"
- using linear_setsum [OF \<open>linear g\<close>] x by auto
+ using linear_sum [OF \<open>linear g\<close>] x by auto
also have "... = (\<Sum>v \<in> C. a v *\<^sub>R g v)"
- using \<open>linear g\<close> by (simp add: linear_setsum linear.scaleR)
+ using \<open>linear g\<close> by (simp add: linear_sum linear.scaleR)
also have "... = (\<Sum>v \<in> C. a v *\<^sub>R gb v)"
- by (simp add: ggb cong: setsum.cong)
+ by (simp add: ggb cong: sum.cong)
finally have "f (g x) = f (\<Sum>v\<in>C. a v *\<^sub>R gb v)" by simp
also have "... = (\<Sum>v\<in>C. f (a v *\<^sub>R gb v))"
- using \<open>linear f\<close> by (simp add: linear_setsum linear.scaleR)
+ using \<open>linear f\<close> by (simp add: linear_sum linear.scaleR)
also have "... = (\<Sum>v\<in>C. a v *\<^sub>R f (gb v))"
by (simp add: \<open>linear f\<close> linear.scaleR)
also have "... = (\<Sum>v\<in>C. a v *\<^sub>R v)"
--- a/src/HOL/Analysis/Poly_Roots.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Poly_Roots.thy Mon Oct 17 14:37:32 2016 +0200
@@ -10,67 +10,67 @@
subsection\<open>Geometric progressions\<close>
-lemma setsum_gp_basic:
+lemma sum_gp_basic:
fixes x :: "'a::{comm_ring,monoid_mult}"
shows "(1 - x) * (\<Sum>i\<le>n. x^i) = 1 - x^Suc n"
by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost)
-lemma setsum_gp0:
+lemma sum_gp0:
fixes x :: "'a::{comm_ring,division_ring}"
shows "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))"
- using setsum_gp_basic[of x n]
+ using sum_gp_basic[of x n]
by (simp add: mult.commute divide_simps)
-lemma setsum_power_add:
+lemma sum_power_add:
fixes x :: "'a::{comm_ring,monoid_mult}"
shows "(\<Sum>i\<in>I. x^(m+i)) = x^m * (\<Sum>i\<in>I. x^i)"
- by (simp add: setsum_distrib_left power_add)
+ by (simp add: sum_distrib_left power_add)
-lemma setsum_power_shift:
+lemma sum_power_shift:
fixes x :: "'a::{comm_ring,monoid_mult}"
assumes "m \<le> n"
shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
proof -
have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
- by (simp add: setsum_distrib_left power_add [symmetric])
+ by (simp add: sum_distrib_left power_add [symmetric])
also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)"
- using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
+ using \<open>m \<le> n\<close> by (intro sum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
finally show ?thesis .
qed
-lemma setsum_gp_multiplied:
+lemma sum_gp_multiplied:
fixes x :: "'a::{comm_ring,monoid_mult}"
assumes "m \<le> n"
shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n"
proof -
have "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)"
- by (metis mult.assoc mult.commute assms setsum_power_shift)
+ by (metis mult.assoc mult.commute assms sum_power_shift)
also have "... =x^m * (1 - x^Suc(n-m))"
- by (metis mult.assoc setsum_gp_basic)
+ by (metis mult.assoc sum_gp_basic)
also have "... = x^m - x^Suc n"
using assms
by (simp add: algebra_simps) (metis le_add_diff_inverse power_add)
finally show ?thesis .
qed
-lemma setsum_gp:
+lemma sum_gp:
fixes x :: "'a::{comm_ring,division_ring}"
shows "(\<Sum>i=m..n. x^i) =
(if n < m then 0
else if x = 1 then of_nat((n + 1) - m)
else (x^m - x^Suc n) / (1 - x))"
-using setsum_gp_multiplied [of m n x]
+using sum_gp_multiplied [of m n x]
apply auto
by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq)
-lemma setsum_gp_offset:
+lemma sum_gp_offset:
fixes x :: "'a::{comm_ring,division_ring}"
shows "(\<Sum>i=m..m+n. x^i) =
(if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
- using setsum_gp [of x m "m+n"]
+ using sum_gp [of x m "m+n"]
by (auto simp: power_add algebra_simps)
-lemma setsum_gp_strict:
+lemma sum_gp_strict:
fixes x :: "'a::{comm_ring,division_ring}"
shows "(\<Sum>i<n. x^i) = (if x = 1 then of_nat n else (1 - x^n) / (1 - x))"
by (induct n) (auto simp: algebra_simps divide_simps)
@@ -84,13 +84,13 @@
proof -
have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
(\<Sum>i\<le>n. a i * (x^i - y^i))"
- by (simp add: algebra_simps setsum_subtractf [symmetric])
+ by (simp add: algebra_simps sum_subtractf [symmetric])
also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
by (simp add: power_diff_sumr2 ac_simps)
also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
- by (simp add: setsum_distrib_left ac_simps)
+ by (simp add: sum_distrib_left ac_simps)
also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
- by (simp add: nested_setsum_swap')
+ by (simp add: nested_sum_swap')
finally show ?thesis .
qed
@@ -102,7 +102,7 @@
{ fix j
have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =
(\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"
- by (rule setsum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto }
+ by (rule sum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto }
then show ?thesis
by (simp add: sub_polyfun)
qed
@@ -115,7 +115,7 @@
{ fix z
have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) =
(z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
- by (simp add: sub_polyfun setsum_distrib_right)
+ by (simp add: sub_polyfun sum_distrib_right)
then have "(\<Sum>i\<le>n. c i * z^i) =
(z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
+ (\<Sum>i\<le>n. c i * a^i)"
@@ -241,7 +241,7 @@
have "\<exists>k\<le>n. b k \<noteq> 0"
apply (rule ccontr)
using polyfun_extremal [OF extr_prem, of 1]
- apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc)
+ apply (auto simp: eventually_at_infinity b simp del: sum_atMost_Suc)
apply (drule_tac x="of_real ba" in spec, simp)
done
then show ?thesis using Suc.IH [of b] ins_ab
--- a/src/HOL/Analysis/Polytope.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Polytope.thy Mon Oct 17 14:37:32 2016 +0200
@@ -293,8 +293,8 @@
proof -
have waff: "w \<in> affine hull T"
using convex_hull_subset_affine_hull w by blast
- obtain a b where a: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> a i" and asum: "setsum a S = 1" and aeqx: "(\<Sum>i\<in>S. a i *\<^sub>R i) = x"
- and b: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> b i" and bsum: "setsum b S = 1" and beqy: "(\<Sum>i\<in>S. b i *\<^sub>R i) = y"
+ obtain a b where a: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> a i" and asum: "sum a S = 1" and aeqx: "(\<Sum>i\<in>S. a i *\<^sub>R i) = x"
+ and b: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> b i" and bsum: "sum b S = 1" and beqy: "(\<Sum>i\<in>S. b i *\<^sub>R i) = y"
using x y by (auto simp: assms convex_hull_finite)
obtain u where "(1 - u) *\<^sub>R x + u *\<^sub>R y \<in> convex hull T" "x \<noteq> y" and weq: "w = (1 - u) *\<^sub>R x + u *\<^sub>R y"
and u01: "0 < u" "u < 1"
@@ -302,43 +302,43 @@
define c where "c i = (1 - u) * a i + u * b i" for i
have cge0: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> c i"
using a b u01 by (simp add: c_def)
- have sumc1: "setsum c S = 1"
- by (simp add: c_def setsum.distrib setsum_distrib_left [symmetric] asum bsum)
+ have sumc1: "sum c S = 1"
+ by (simp add: c_def sum.distrib sum_distrib_left [symmetric] asum bsum)
have sumci_xy: "(\<Sum>i\<in>S. c i *\<^sub>R i) = (1 - u) *\<^sub>R x + u *\<^sub>R y"
- apply (simp add: c_def setsum.distrib scaleR_left_distrib)
- by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric] aeqx beqy)
+ apply (simp add: c_def sum.distrib scaleR_left_distrib)
+ by (simp only: scaleR_scaleR [symmetric] Real_Vector_Spaces.scaleR_right.sum [symmetric] aeqx beqy)
show ?thesis
- proof (cases "setsum c (S - T) = 0")
+ proof (cases "sum c (S - T) = 0")
case True
have ci0: "\<And>i. i \<in> (S - T) \<Longrightarrow> c i = 0"
- using True cge0 by (simp add: \<open>finite S\<close> setsum_nonneg_eq_0_iff)
+ using True cge0 by (simp add: \<open>finite S\<close> sum_nonneg_eq_0_iff)
have a0: "a i = 0" if "i \<in> (S - T)" for i
using ci0 [OF that] u01 a [of i] b [of i] that
by (simp add: c_def Groups.ordered_comm_monoid_add_class.add_nonneg_eq_0_iff)
- have [simp]: "setsum a T = 1"
- using assms by (metis setsum.mono_neutral_cong_right a0 asum)
+ have [simp]: "sum a T = 1"
+ using assms by (metis sum.mono_neutral_cong_right a0 asum)
show ?thesis
apply (simp add: convex_hull_finite \<open>finite T\<close>)
apply (rule_tac x=a in exI)
using a0 assms
- apply (auto simp: cge0 a aeqx [symmetric] setsum.mono_neutral_right)
+ apply (auto simp: cge0 a aeqx [symmetric] sum.mono_neutral_right)
done
next
case False
- define k where "k = setsum c (S - T)"
+ define k where "k = sum c (S - T)"
have "k > 0" using False
- unfolding k_def by (metis DiffD1 antisym_conv cge0 setsum_nonneg not_less)
- have weq_sumsum: "w = setsum (\<lambda>x. c x *\<^sub>R x) T + setsum (\<lambda>x. c x *\<^sub>R x) (S - T)"
- by (metis (no_types) add.commute S(1) S(2) setsum.subset_diff sumci_xy weq)
+ unfolding k_def by (metis DiffD1 antisym_conv cge0 sum_nonneg not_less)
+ have weq_sumsum: "w = sum (\<lambda>x. c x *\<^sub>R x) T + sum (\<lambda>x. c x *\<^sub>R x) (S - T)"
+ by (metis (no_types) add.commute S(1) S(2) sum.subset_diff sumci_xy weq)
show ?thesis
proof (cases "k = 1")
case True
- then have "setsum c T = 0"
- by (simp add: S k_def setsum_diff sumc1)
- then have [simp]: "setsum c (S - T) = 1"
- by (simp add: S setsum_diff sumc1)
+ then have "sum c T = 0"
+ by (simp add: S k_def sum_diff sumc1)
+ then have [simp]: "sum c (S - T) = 1"
+ by (simp add: S sum_diff sumc1)
have ci0: "\<And>i. i \<in> T \<Longrightarrow> c i = 0"
- by (meson \<open>finite T\<close> \<open>setsum c T = 0\<close> \<open>T \<subseteq> S\<close> cge0 setsum_nonneg_eq_0_iff subsetCE)
+ by (meson \<open>finite T\<close> \<open>sum c T = 0\<close> \<open>T \<subseteq> S\<close> cge0 sum_nonneg_eq_0_iff subsetCE)
then have [simp]: "(\<Sum>i\<in>S-T. c i *\<^sub>R i) = w"
by (simp add: weq_sumsum)
have "w \<in> convex hull (S - T)"
@@ -350,22 +350,22 @@
using disj waff by blast
next
case False
- then have sumcf: "setsum c T = 1 - k"
- by (simp add: S k_def setsum_diff sumc1)
+ then have sumcf: "sum c T = 1 - k"
+ by (simp add: S k_def sum_diff sumc1)
have "(\<Sum>i\<in>T. c i *\<^sub>R i) /\<^sub>R (1 - k) \<in> convex hull T"
apply (simp add: convex_hull_finite fin)
apply (rule_tac x="\<lambda>i. inverse (1-k) * c i" in exI)
apply auto
- apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) setsum_nonneg subsetCE)
- apply (metis False mult.commute right_inverse right_minus_eq setsum_distrib_left sumcf)
- by (metis (mono_tags, lifting) scaleR_right.setsum scaleR_scaleR setsum.cong)
- with \<open>0 < k\<close> have "inverse(k) *\<^sub>R (w - setsum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
+ apply (metis sumcf cge0 inverse_nonnegative_iff_nonnegative mult_nonneg_nonneg S(2) sum_nonneg subsetCE)
+ apply (metis False mult.commute right_inverse right_minus_eq sum_distrib_left sumcf)
+ by (metis (mono_tags, lifting) scaleR_right.sum scaleR_scaleR sum.cong)
+ with \<open>0 < k\<close> have "inverse(k) *\<^sub>R (w - sum (\<lambda>i. c i *\<^sub>R i) T) \<in> affine hull T"
by (simp add: affine_diff_divide [OF affine_affine_hull] False waff convex_hull_subset_affine_hull [THEN subsetD])
- moreover have "inverse(k) *\<^sub>R (w - setsum (\<lambda>x. c x *\<^sub>R x) T) \<in> convex hull (S - T)"
+ moreover have "inverse(k) *\<^sub>R (w - sum (\<lambda>x. c x *\<^sub>R x) T) \<in> convex hull (S - T)"
apply (simp add: weq_sumsum convex_hull_finite fin)
apply (rule_tac x="\<lambda>i. inverse k * c i" in exI)
using \<open>k > 0\<close> cge0
- apply (auto simp: scaleR_right.setsum setsum_distrib_left [symmetric] k_def [symmetric])
+ apply (auto simp: scaleR_right.sum sum_distrib_left [symmetric] k_def [symmetric])
done
ultimately show ?thesis
using disj by blast
@@ -2727,7 +2727,7 @@
have "norm (x - y) \<le> (\<Sum>b\<in>Basis. \<bar>(x-y) \<bullet> b\<bar>)"
by (rule norm_le_l1)
also have "... \<le> of_nat (DIM('a)) * (e / 2 / DIM('a))"
- proof (rule setsum_bounded_above)
+ proof (rule sum_bounded_above)
fix i::'a
assume "i \<in> Basis"
then have I': "\<And>z b. \<lbrakk>z \<in> C; b = z * e / (2 * real DIM('a))\<rbrakk> \<Longrightarrow> i \<bullet> x \<le> b \<and> i \<bullet> y \<le> b \<or> i \<bullet> x \<ge> b \<and> i \<bullet> y \<ge> b"
--- a/src/HOL/Analysis/Regularity.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Regularity.thy Mon Oct 17 14:37:32 2016 +0200
@@ -260,7 +260,7 @@
also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
by (intro summable_LIMSEQ) auto
finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
- by (simp add: emeasure_eq_measure measure_nonneg setsum_nonneg)
+ by (simp add: emeasure_eq_measure measure_nonneg sum_nonneg)
have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
case 1
@@ -274,12 +274,12 @@
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 "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
- by (auto simp add: emeasure_eq_measure setsum_nonneg measure_nonneg)
- also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule setsum_le_suminf) auto
+ by (auto simp add: emeasure_eq_measure sum_nonneg measure_nonneg)
+ also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule sum_le_suminf) auto
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<n0. measure M (D i)) < e/2"
- using n0 by (auto simp: measure_nonneg setsum_nonneg)
+ using n0 by (auto simp: measure_nonneg sum_nonneg)
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
fix i
@@ -301,16 +301,16 @@
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 \<open>0 < e\<close>
- by (auto intro: setsum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
+ by (auto intro: sum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
- by (simp add: setsum.distrib)
+ by (simp add: sum.distrib)
also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) + e / 2" using \<open>0 < e\<close>
by (auto simp: field_simps intro!: mult_left_mono)
finally
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"
- using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
+ using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg sum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
moreover
have "?K \<subseteq> (\<Union>i. D i)" using K by auto
moreover
@@ -332,7 +332,7 @@
from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
using \<open>0<e\<close>
- by (auto simp: emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
+ by (auto simp: emeasure_eq_measure measure_nonneg sum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
finite_measure_mono sb
simp del: ennreal_plus)
qed
--- a/src/HOL/Analysis/Summation_Tests.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Summation_Tests.thy Mon Oct 17 14:37:32 2016 +0200
@@ -218,7 +218,7 @@
assumes mono: "\<And>m n. 0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> f n \<le> f m"
shows "(\<Sum>k=1..<n. f k) \<ge> (\<Sum>k=1..<n. f (2 * 2 ^ natlog2 k))" (is "?thesis1")
"(\<Sum>k=1..<n. f k) \<le> (\<Sum>k=1..<n. f (2 ^ natlog2 k))" (is "?thesis2")
- by (intro setsum_mono mono pow_natlog2_ge pow_natlog2_le, simp, simp)+
+ by (intro sum_mono mono pow_natlog2_ge pow_natlog2_le, simp, simp)+
private lemma condensation_condense1: "(\<Sum>k=1..<2^n. f (2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ k))"
proof (induction n)
@@ -226,10 +226,10 @@
have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
also have "(\<Sum>k\<in>\<dots>. f (2 ^ natlog2 k)) =
(\<Sum>k<n. 2^k * f (2^k)) + (\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k))"
- by (subst setsum.union_disjoint) (insert Suc, auto)
+ by (subst sum.union_disjoint) (insert Suc, auto)
also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
hence "(\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
- by (intro setsum.cong) simp_all
+ by (intro sum.cong) simp_all
also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
finally show ?case by simp
qed simp
@@ -240,10 +240,10 @@
have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
also have "(\<Sum>k\<in>\<dots>. f (2 * 2 ^ natlog2 k)) =
(\<Sum>k<n. 2^k * f (2^Suc k)) + (\<Sum>k = 2^n..<2^Suc n. f (2 * 2^natlog2 k))"
- by (subst setsum.union_disjoint) (insert Suc, auto)
+ by (subst sum.union_disjoint) (insert Suc, auto)
also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
- by (intro setsum.cong) simp_all
+ by (intro sum.cong) simp_all
also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
finally show ?case by simp
qed simp
@@ -267,11 +267,11 @@
hence "convergent (\<lambda>n. \<Sum>k<n. f' k) \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. f' k)"
by (rule monoseq_imp_convergent_iff_Bseq)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f' k)" unfolding One_nat_def
- by (subst setsum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
+ by (subst sum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f k)" unfolding f'_def by simp
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
by (rule nonneg_incseq_Bseq_subseq_iff[symmetric])
- (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def)
+ (auto intro!: sum_nonneg incseq_SucI nonneg simp: subseq_def)
also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^k))"
proof (intro iffI)
assume A: "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
@@ -279,20 +279,20 @@
proof (intro always_eventually allI)
fix n :: nat
have "norm (\<Sum>k<n. 2^k * f (2^Suc k)) = (\<Sum>k<n. 2^k * f (2^Suc k))" unfolding real_norm_def
- by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+ by (intro abs_of_nonneg sum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
also have "\<dots> \<le> (\<Sum>k=1..<2^n. f k)"
by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono')
also have "\<dots> = norm \<dots>" unfolding real_norm_def
- by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+ by (intro abs_of_nonneg[symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg)
finally show "norm (\<Sum>k<n. 2 ^ k * f (2 ^ Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)" .
qed
from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
- by (simp add: setsum_distrib_left setsum_distrib_right mult_ac)
+ by (simp add: sum_distrib_left sum_distrib_right mult_ac)
hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
- by (intro Bseq_add, subst setsum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
+ by (intro Bseq_add, subst sum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
- by (subst setsum_head_upt_Suc) (simp_all add: add_ac)
+ by (subst sum_head_upt_Suc) (simp_all add: add_ac)
thus "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
by (subst (asm) Bseq_Suc_iff) (simp add: atLeast0LessThan)
next
@@ -301,11 +301,11 @@
proof (intro always_eventually allI)
fix n :: nat
have "norm (\<Sum>k=1..<2^n. f k) = (\<Sum>k=1..<2^n. f k)" unfolding real_norm_def
- by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+ by (intro abs_of_nonneg sum_nonneg ballI mult_nonneg_nonneg nonneg)
also have "\<dots> \<le> (\<Sum>k<n. 2^k * f (2^k))"
by (subst condensation_condense1 [symmetric]) (intro condensation_inequality mono')
also have "\<dots> = norm \<dots>" unfolding real_norm_def
- by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+ by (intro abs_of_nonneg [symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
finally show "norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))" .
qed
from this and A show "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)" by (rule Bseq_eventually_mono)
@@ -384,7 +384,7 @@
by (simp add: summable_iff_convergent convergent_norm)
hence "convergent (\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real)" by (simp only: norm_of_real)
also have "(\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real) = (\<lambda>n. \<Sum>k<n. inverse (of_nat k))"
- by (intro ext abs_of_nonneg setsum_nonneg) auto
+ by (intro ext abs_of_nonneg sum_nonneg) auto
also have "convergent \<dots> \<longleftrightarrow> summable (\<lambda>k. inverse (of_nat k) :: real)"
by (simp add: summable_iff_convergent)
finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus)
@@ -424,31 +424,31 @@
have n: "n > m" by (simp add: n_def)
from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
- by (simp add: setsum_distrib_left[symmetric] abs_mult)
+ by (simp add: sum_distrib_left[symmetric] abs_mult)
also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
- by (subst setsum.union_disjoint) auto
+ by (subst sum.union_disjoint) auto
also have "norm \<dots> \<le> norm (\<Sum>k\<le>m. r * f k) + norm (\<Sum>k=Suc m..n. r * f k)"
by (rule norm_triangle_ineq)
also from r less_imp_le[OF m(1)] have "(\<Sum>k=Suc m..n. r * f k) \<ge> 0"
- by (intro setsum_nonneg) auto
+ by (intro sum_nonneg) auto
hence "norm (\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=Suc m..n. r * f k)" by simp
also have "(\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=m..<n. r * f (Suc k))"
- by (subst setsum_shift_bounds_Suc_ivl [symmetric])
+ by (subst sum_shift_bounds_Suc_ivl [symmetric])
(simp only: atLeastLessThanSuc_atLeastAtMost)
also from m have "\<dots> \<le> (\<Sum>k=m..<n. p k * f k - p (Suc k) * f (Suc k))"
- by (intro setsum_mono[OF less_imp_le]) simp_all
+ by (intro sum_mono[OF less_imp_le]) simp_all
also have "\<dots> = -(\<Sum>k=m..<n. p (Suc k) * f (Suc k) - p k * f k)"
- by (simp add: setsum_negf [symmetric] algebra_simps)
+ by (simp add: sum_negf [symmetric] algebra_simps)
also from n have "\<dots> = p m * f m - p n * f n"
- by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all
+ by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst sum_Suc_diff) simp_all
also from less_imp_le[OF m(1)] m(2) n have "\<dots> \<le> p m * f m" by simp
finally show "norm (\<Sum>k\<le>n. f k) \<le> (norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r" using r
by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac)
qed
moreover have "(\<Sum>k\<le>n. f k) \<le> (\<Sum>k\<le>n'. f k)" if "Suc m \<le> n" "n \<le> n'" for n n'
- using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto
+ using less_imp_le[OF m(1)] that by (intro sum_mono2) auto
ultimately show "convergent (\<lambda>n. \<Sum>k\<le>n. f k)" by (rule Bseq_monoseq_convergent'_inc)
qed
@@ -905,7 +905,7 @@
from r have "summable (\<lambda>n. (\<Sum>i\<le>n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))"
by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all
thus "summable (\<lambda>n. (\<Sum>i\<le>n. f i * g (n - i)) * of_real r ^ n)"
- by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_setsum_right)
+ by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_sum_right)
qed
lemma le_conv_radius_iff:
--- a/src/HOL/Analysis/Tagged_Division.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Tagged_Division.thy Mon Oct 17 14:37:32 2016 +0200
@@ -30,8 +30,8 @@
lemma sum_sum_product:
assumes "finite s"
and "\<forall>i\<in>s. finite (t i)"
- shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s =
- setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
+ shows "sum (\<lambda>i. sum (x i) (t i)::real) s =
+ sum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
using assms
proof induct
case (insert a s)
@@ -39,14 +39,14 @@
(\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
show ?case
unfolding *
- apply (subst setsum.union_disjoint)
- unfolding setsum.insert[OF insert(1-2)]
+ apply (subst sum.union_disjoint)
+ unfolding sum.insert[OF insert(1-2)]
prefer 4
apply (subst insert(3))
unfolding add_right_cancel
proof -
- show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
- apply (subst setsum.reindex)
+ show "sum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
+ apply (subst sum.reindex)
unfolding inj_on_def
apply auto
done
@@ -196,13 +196,13 @@
lemma interval_upperbound[simp]:
"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
- unfolding interval_upperbound_def euclidean_representation_setsum cbox_def
+ unfolding interval_upperbound_def euclidean_representation_sum cbox_def
by (safe intro!: cSup_eq) auto
lemma interval_lowerbound[simp]:
"\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
- unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def
+ unfolding interval_lowerbound_def euclidean_representation_sum cbox_def
by (safe intro!: cInf_eq) auto
lemmas interval_bounds = interval_upperbound interval_lowerbound
@@ -230,7 +230,7 @@
have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
ultimately show ?thesis unfolding interval_upperbound_def
- by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+ by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
qed
lemma interval_lowerbound_Times:
@@ -244,7 +244,7 @@
have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
ultimately show ?thesis unfolding interval_lowerbound_def
- by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+ by (subst sum_Basis_prod_eq) (auto simp add: sum_prod)
qed
subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
@@ -1359,7 +1359,7 @@
assumes g: "operative g" shows "g {} = \<^bold>1"
proof -
have *: "cbox One (-One) = ({}::'b set)"
- by (auto simp: box_eq_empty inner_setsum_left inner_Basis setsum.If_cases ex_in_conv)
+ by (auto simp: box_eq_empty inner_sum_left inner_Basis sum.If_cases ex_in_conv)
moreover have "box One (-One) = ({}::'b set)"
using box_subset_cbox[of One "-One"] by (auto simp: *)
ultimately show ?thesis
@@ -1438,7 +1438,7 @@
apply rule
unfolding mem_Collect_eq split_beta
apply (erule bexE conjE)+
- apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
+ apply (simp only: mem_Collect_eq inner_sum_left_Basis simp_thms)
apply (erule exE conjE)+
proof
fix i l x
@@ -1931,7 +1931,7 @@
fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "a \<le> b"
and "p tagged_division_of {a..b}"
- shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
+ shows "sum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
proof -
let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
@@ -1941,11 +1941,11 @@
by auto
have **: "cbox a b \<noteq> {}"
using assms(1) by auto
- note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
+ note sum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric]
show ?thesis
unfolding *
- apply (rule setsum.cong)
+ apply (rule sum.cong)
unfolding split_paired_all split_conv
using assms(2)
apply auto
@@ -1959,7 +1959,7 @@
fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
assumes "a \<le> b"
and "p tagged_division_of {a..b}"
- shows "setsum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
+ shows "sum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
using additive_tagged_division_1[OF _ assms(2), of f]
using assms(1)
by auto
@@ -2264,24 +2264,24 @@
if e: "0 < e" for e
proof -
obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
- using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
+ using real_arch_pow[of 2 "(sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
show ?thesis
proof (rule exI [where x=n], clarify)
fix x y
assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
- have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
+ have "dist x y \<le> sum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
unfolding dist_norm by(rule norm_le_l1)
- also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
- proof (rule setsum_mono)
+ also have "\<dots> \<le> sum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
+ proof (rule sum_mono)
fix i :: 'a
assume i: "i \<in> Basis"
show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
using xy[unfolded mem_box,THEN bspec, OF i]
by (auto simp: inner_diff_left)
qed
- also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
- unfolding setsum_divide_distrib
- proof (rule setsum_mono)
+ also have "\<dots> \<le> sum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
+ unfolding sum_divide_distrib
+ proof (rule sum_mono)
show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
proof (induct n)
case 0
@@ -2600,9 +2600,9 @@
if "(\<And>i. i \<in> Basis \<Longrightarrow> \<bar>x \<bullet> i - y \<bullet> i\<bar> \<le> \<epsilon> / (2 * real DIM('a)))" for y
proof -
have "norm (x - y) \<le> (\<Sum>i\<in>Basis. \<bar>x \<bullet> i - y \<bullet> i\<bar>)"
- by (metis (no_types, lifting) inner_diff_left norm_le_l1 setsum.cong)
+ by (metis (no_types, lifting) inner_diff_left norm_le_l1 sum.cong)
also have "... \<le> DIM('a) * (\<epsilon> / (2 * real DIM('a)))"
- by (meson setsum_bounded_above that)
+ by (meson sum_bounded_above that)
also have "... = \<epsilon> / 2"
by (simp add: divide_simps)
also have "... < \<epsilon>"
@@ -2814,4 +2814,4 @@
"eventually (\<lambda>p. p tagged_division_of s) (division_filter s)"
unfolding eventually_division_filter by (intro exI[of _ "\<lambda>x. ball x 1"]) auto
-end
\ No newline at end of file
+end
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Mon Oct 17 14:37:32 2016 +0200
@@ -58,22 +58,22 @@
lemma finite_support[intro]: "finite s \<Longrightarrow> finite (support_on s f)"
unfolding support_on_def by auto
-(* TODO: is supp_setsum really needed? TODO: Generalize to Finite_Set.fold *)
-definition (in comm_monoid_add) supp_setsum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
+(* TODO: is supp_sum really needed? TODO: Generalize to Finite_Set.fold *)
+definition (in comm_monoid_add) supp_sum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
where
- "supp_setsum f s = (\<Sum>x\<in>support_on s f. f x)"
-
-lemma supp_setsum_empty[simp]: "supp_setsum f {} = 0"
- unfolding supp_setsum_def by auto
-
-lemma supp_setsum_insert[simp]:
+ "supp_sum f s = (\<Sum>x\<in>support_on s f. f x)"
+
+lemma supp_sum_empty[simp]: "supp_sum f {} = 0"
+ unfolding supp_sum_def by auto
+
+lemma supp_sum_insert[simp]:
"finite (support_on s f) \<Longrightarrow>
- supp_setsum f (insert x s) = (if x \<in> s then supp_setsum f s else f x + supp_setsum f s)"
- by (simp add: supp_setsum_def in_support_on insert_absorb)
-
-lemma supp_setsum_divide_distrib: "supp_setsum f A / (r::'a::field) = supp_setsum (\<lambda>n. f n / r) A"
+ supp_sum f (insert x s) = (if x \<in> s then supp_sum f s else f x + supp_sum f s)"
+ by (simp add: supp_sum_def in_support_on insert_absorb)
+
+lemma supp_sum_divide_distrib: "supp_sum f A / (r::'a::field) = supp_sum (\<lambda>n. f n / r) A"
by (cases "r = 0")
- (auto simp: supp_setsum_def setsum_divide_distrib intro!: setsum.cong support_on_cong)
+ (auto simp: supp_sum_def sum_divide_distrib intro!: sum.cong support_on_cong)
(*END OF SUPPORT, ETC.*)
@@ -1164,7 +1164,7 @@
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)"
unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
- proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
+ proof (rule real_sqrt_less_mono, rule sum_strict_mono)
fix i :: "'a"
assume i: "i \<in> Basis"
have "a i < y\<bullet>i \<and> y\<bullet>i < b i"
@@ -1497,7 +1497,7 @@
moreover have "\<And>x b::'a. \<And>n::nat. \<bar>x \<bullet> b\<bar> < real n \<longleftrightarrow> - real n < x \<bullet> b \<and> x \<bullet> b < real n"
by auto
ultimately show ?thesis
- by (auto simp: box_def inner_setsum_left inner_Basis setsum.If_cases)
+ by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases)
qed
text \<open>Intervals in general, including infinite and mixtures of open and closed.\<close>
@@ -1562,10 +1562,10 @@
qed
lemma cbox01_nonempty [simp]: "cbox 0 One \<noteq> {}"
- by (simp add: box_ne_empty inner_Basis inner_setsum_left setsum_nonneg)
+ by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg)
lemma box01_nonempty [simp]: "box 0 One \<noteq> {}"
- by (simp add: box_ne_empty inner_Basis inner_setsum_left) (simp add: setsum.remove)
+ by (simp add: box_ne_empty inner_Basis inner_sum_left) (simp add: sum.remove)
subsection\<open>Connectedness\<close>
@@ -5046,10 +5046,10 @@
have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))"
apply (subst euclidean_dist_l2)
using zero_le_dist
- apply (rule setL2_le_setsum)
+ apply (rule setL2_le_sum)
done
also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
- apply (rule setsum_strict_mono)
+ apply (rule sum_strict_mono)
using n
apply auto
done
@@ -5448,7 +5448,7 @@
by (frule summable_LIMSEQ_zero) (simp add: convergent_imp_bounded)
lemma summable_imp_sums_bounded:
- "summable f \<Longrightarrow> bounded (range (\<lambda>n. setsum f {..<n}))"
+ "summable f \<Longrightarrow> bounded (range (\<lambda>n. sum f {..<n}))"
by (auto simp: summable_def sums_def dest: convergent_imp_bounded)
lemma power_series_conv_imp_absconv_weak:
@@ -7064,17 +7064,17 @@
show ?thesis
proof (rule image_eqI)
show "y = f (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i))"
- apply (simp add: linear_add linear_setsum linear.scaleR \<open>linear f\<close> surj_f_inv_f \<open>surj f\<close>)
+ apply (simp add: linear_add linear_sum linear.scaleR \<open>linear f\<close> surj_f_inv_f \<open>surj f\<close>)
apply (simp add: euclidean_representation u_def)
done
have "dist (x + (\<Sum>i\<in>Basis. (u \<bullet> i) *\<^sub>R inv f i)) x \<le> (\<Sum>i\<in>Basis. norm ((u \<bullet> i) *\<^sub>R inv f i))"
- by (simp add: dist_norm setsum_norm_le)
+ by (simp add: dist_norm sum_norm_le)
also have "... = (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar> * norm (inv f i))"
by (simp add: )
also have "... \<le> (\<Sum>i\<in>Basis. \<bar>u \<bullet> i\<bar>) * B"
- by (simp add: B setsum_distrib_right setsum_mono mult_left_mono)
+ by (simp add: B sum_distrib_right sum_mono mult_left_mono)
also have "... \<le> DIM('b) * dist y (f x) * B"
- apply (rule mult_right_mono [OF setsum_bounded_above])
+ apply (rule mult_right_mono [OF sum_bounded_above])
using \<open>0 < B\<close> by (auto simp add: Basis_le_norm dist_norm u_def)
also have "... < e"
by (metis mult.commute mult.left_commute that)
@@ -8200,7 +8200,7 @@
}
then have "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b"
apply -
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply auto
done
then have "norm x \<le> ?b"
@@ -9164,9 +9164,9 @@
have "finite d"
using finite_subset [OF d finite_Basis] .
then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d"
- by (simp add: span_setsum span_clauses)
+ by (simp add: span_sum span_clauses)
also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)"
- by (rule setsum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp add: x)
+ by (rule sum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp add: x)
finally show "x \<in> span d"
unfolding euclidean_representation .
qed
@@ -9950,10 +9950,10 @@
have *: "setL2 (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis < e"
if "\<forall>i\<in>Basis. dist (f x \<bullet> i) (l \<bullet> i) < e / real DIM('b)" for x
proof -
- have "setL2 (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> setsum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis"
- by (simp add: setL2_le_setsum)
+ have "setL2 (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis \<le> sum (\<lambda>i. dist (f x \<bullet> i) (l \<bullet> i)) Basis"
+ by (simp add: setL2_le_sum)
also have "... < DIM('b) * (e / real DIM('b))"
- apply (rule setsum_bounded_above_strict)
+ apply (rule sum_bounded_above_strict)
using that by auto
also have "... = e"
by (simp add: field_simps)
@@ -9998,14 +9998,14 @@
by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib)
then obtain F where F: "\<And>i x. i \<in> Basis \<Longrightarrow> \<bar>f' x \<bullet> i\<bar> \<le> norm x * F i"
by metis
- have "norm (f' x) \<le> norm x * setsum F Basis" for x
+ have "norm (f' x) \<le> norm x * sum F Basis" for x
proof -
have "norm (f' x) \<le> (\<Sum>i\<in>Basis. \<bar>f' x \<bullet> i\<bar>)"
by (rule norm_le_l1)
also have "... \<le> (\<Sum>i\<in>Basis. norm x * F i)"
- by (metis F setsum_mono)
- also have "... = norm x * setsum F Basis"
- by (simp add: setsum_distrib_left)
+ by (metis F sum_mono)
+ also have "... = norm x * sum F Basis"
+ by (simp add: sum_distrib_left)
finally show ?thesis .
qed
then show ?lhs
--- a/src/HOL/Analysis/Uniform_Limit.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Uniform_Limit.thy Mon Oct 17 14:37:32 2016 +0200
@@ -290,7 +290,7 @@
corollary series_comparison_uniform:
fixes f :: "_ \<Rightarrow> nat \<Rightarrow> _ :: banach"
assumes g: "summable g" and le: "\<And>n x. N \<le> n \<and> x \<in> A \<Longrightarrow> norm(f x n) \<le> g n"
- shows "\<exists>l. \<forall>e. 0 < e \<longrightarrow> (\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> A \<longrightarrow> dist(setsum (f x) {..<n}) (l x) < e)"
+ shows "\<exists>l. \<forall>e. 0 < e \<longrightarrow> (\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> A \<longrightarrow> dist(sum (f x) {..<n}) (l x) < e)"
proof -
have 1: "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>A. norm (f x n) \<le> g n"
using le eventually_sequentially by auto
--- a/src/HOL/Analysis/Weierstrass_Theorems.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/Weierstrass_Theorems.thy Mon Oct 17 14:37:32 2016 +0200
@@ -25,7 +25,7 @@
"(\<Sum>k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)"
apply (rule DERIV_unique [where f = "\<lambda>a. (a+b)^n" and x=a])
apply (subst binomial_ring)
- apply (rule derivative_eq_intros setsum.cong | simp)+
+ apply (rule derivative_eq_intros sum.cong | simp)+
done
lemma binomial_deriv2:
@@ -33,21 +33,21 @@
of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)"
apply (rule DERIV_unique [where f = "\<lambda>a. of_nat n * (a+b::real) ^ (n-1)" and x=a])
apply (subst binomial_deriv1 [symmetric])
- apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+
+ apply (rule derivative_eq_intros sum.cong | simp add: Num.numeral_2_eq_2)+
done
lemma sum_k_Bernstein [simp]: "(\<Sum>k = 0..n. real k * Bernstein n k x) = of_nat n * x"
apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
- apply (simp add: setsum_distrib_right)
- apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong)
+ apply (simp add: sum_distrib_right)
+ apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: sum.cong)
done
lemma sum_kk_Bernstein [simp]: "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2"
proof -
have "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric])
- apply (simp add: setsum_distrib_right)
- apply (rule setsum.cong [OF refl])
+ apply (simp add: sum_distrib_right)
+ apply (rule sum.cong [OF refl])
apply (simp add: Bernstein_def power2_eq_square algebra_simps)
apply (rename_tac k)
apply (subgoal_tac "k = 0 \<or> k = 1 \<or> (\<exists>k'. k = Suc (Suc k'))")
@@ -97,14 +97,14 @@
have *: "\<And>a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x"
by (simp add: algebra_simps power2_eq_square)
have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
- apply (simp add: * setsum.distrib)
- apply (simp add: setsum_distrib_left [symmetric] mult.assoc)
+ apply (simp add: * sum.distrib)
+ apply (simp add: sum_distrib_left [symmetric] mult.assoc)
apply (simp add: algebra_simps power2_eq_square)
done
then have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
by (simp add: power2_eq_square)
then show ?thesis
- using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute)
+ using n by (simp add: sum_divide_distrib divide_simps mult.commute power2_commute)
qed
{ fix k
assume k: "k \<le> n"
@@ -138,14 +138,14 @@
qed
} note * = this
have "\<bar>f x - (\<Sum> k = 0..n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum> k = 0..n. (f x - f(k / n)) * Bernstein n k x\<bar>"
- by (simp add: setsum_subtractf setsum_distrib_left [symmetric] algebra_simps)
+ by (simp add: sum_subtractf sum_distrib_left [symmetric] algebra_simps)
also have "... \<le> (\<Sum> k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
- apply (rule order_trans [OF setsum_abs setsum_mono])
+ apply (rule order_trans [OF sum_abs sum_mono])
using *
apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
done
also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
- apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_distrib_left [symmetric] mult.assoc sum_bern)
+ apply (simp only: sum.distrib Rings.semiring_class.distrib_right sum_distrib_left [symmetric] mult.assoc sum_bern)
using \<open>d>0\<close> x
apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
done
@@ -189,7 +189,7 @@
lemma power: "f \<in> R \<Longrightarrow> (\<lambda>x. f x ^ n) \<in> R"
by (induct n) (auto simp: const mult)
- lemma setsum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
+ lemma sum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
by (induct I rule: finite_induct; simp add: const add)
lemma setprod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
@@ -268,16 +268,16 @@
by (simp add: card_gt_0_iff)
define p where [abs_def]: "p x = (1 / card subU) * (\<Sum>t \<in> subU. pf t x)" for x
have pR: "p \<in> R"
- unfolding p_def using subU pf by (fast intro: pf const mult setsum)
+ unfolding p_def using subU pf by (fast intro: pf const mult sum)
have pt0 [simp]: "p t0 = 0"
- using subU pf by (auto simp: p_def intro: setsum.neutral)
+ using subU pf by (auto simp: p_def intro: sum.neutral)
have pt_pos: "p t > 0" if t: "t \<in> S-U" for t
proof -
obtain i where i: "i \<in> subU" "t \<in> Uf i" using subU t by blast
show ?thesis
using subU i t
apply (clarsimp simp: p_def divide_simps)
- apply (rule setsum_pos2 [OF \<open>finite subU\<close>])
+ apply (rule sum_pos2 [OF \<open>finite subU\<close>])
using Uf t pf01 apply auto
apply (force elim!: subsetCE)
done
@@ -286,13 +286,13 @@
proof -
have "0 \<le> p x"
using subU cardp t
- apply (simp add: p_def divide_simps setsum_nonneg)
- apply (rule setsum_nonneg)
+ apply (simp add: p_def divide_simps sum_nonneg)
+ apply (rule sum_nonneg)
using pf01 by force
moreover have "p x \<le> 1"
using subU cardp t
- apply (simp add: p_def divide_simps setsum_nonneg)
- apply (rule setsum_bounded_above [where 'a=real and K=1, simplified])
+ apply (simp add: p_def divide_simps sum_nonneg)
+ apply (rule sum_bounded_above [where 'a=real and K=1, simplified])
using pf01 by force
ultimately show ?thesis
by auto
@@ -449,12 +449,12 @@
by blast
have tVft: "\<And>w. w \<in> A \<Longrightarrow> w \<in> Vf w"
using Vf by blast
- then have setsum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
+ then have sum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
by blast
have com_A: "compact A" using A
by (metis compact compact_Int_closed inf.absorb_iff2)
obtain subA where subA: "subA \<subseteq> A" "finite subA" "A \<subseteq> (\<Union>x \<in> subA. Vf x)"
- by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0])
+ by (blast intro: that open_Vf compactE_image [OF com_A _ sum_max_0])
then have [simp]: "subA \<noteq> {}"
using \<open>a \<in> A\<close> by auto
then have cardp: "card subA > 0" using subA
@@ -473,12 +473,12 @@
proof -
have "0 \<le> pff x"
using subA cardp t
- apply (simp add: pff_def divide_simps setsum_nonneg)
+ apply (simp add: pff_def divide_simps sum_nonneg)
apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg)
using ff by fastforce
moreover have "pff x \<le> 1"
using subA cardp t
- apply (simp add: pff_def divide_simps setsum_nonneg)
+ apply (simp add: pff_def divide_simps sum_nonneg)
apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
using ff by fastforce
ultimately show ?thesis
@@ -594,9 +594,9 @@
by metis
define g where [abs_def]: "g x = e * (\<Sum>i\<le>n. xf i x)" for x
have gR: "g \<in> R"
- unfolding g_def by (fast intro: mult const setsum xfR)
+ unfolding g_def by (fast intro: mult const sum xfR)
have gge0: "\<And>x. x \<in> S \<Longrightarrow> g x \<ge> 0"
- using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
+ using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg)
have A0: "A 0 = {}"
using fpos e by (fastforce simp: A_def)
have An: "A n = S"
@@ -647,14 +647,14 @@
have "g t = e * (\<Sum>i<j. xf i t) + e * (\<Sum>i=j..n. xf i t)"
using j1 jn e
apply (simp add: g_def distrib_left [symmetric])
- apply (subst setsum.union_disjoint [symmetric])
+ apply (subst sum.union_disjoint [symmetric])
apply (auto simp: ivl_disj_un)
done
also have "... \<le> e*j + e * ((Suc n - j)*e/n)"
apply (rule add_mono)
apply (simp_all only: mult_le_cancel_left_pos e)
- apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
- using setsum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
+ apply (rule sum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
+ using sum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
apply simp
done
also have "... \<le> j*e + e*(n - j + 1)*e/n "
@@ -680,15 +680,15 @@
also have "... \<le> e * (\<Sum>i\<le>j-2. xf i t)"
using e
apply simp
- apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]])
+ apply (rule order_trans [OF _ sum_bounded_below [OF less_imp_le [OF ****]]])
using True
apply (simp_all add: of_nat_Suc of_nat_diff)
done
also have "... \<le> g t"
using jn e
using e xf01 t
- apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
- apply (rule Groups_Big.setsum_mono2, auto)
+ apply (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg)
+ apply (rule Groups_Big.sum_mono2, auto)
done
finally show ?thesis .
qed
@@ -848,8 +848,8 @@
unfolding add_uminus_conv_diff [symmetric]
by (metis polynomial_function_add polynomial_function_minus)
-lemma polynomial_function_setsum [intro]:
- "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. setsum (f x) I)"
+lemma polynomial_function_sum [intro]:
+ "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. sum (f x) I)"
by (induct I rule: finite_induct) auto
lemma real_polynomial_function_minus [intro]:
@@ -862,9 +862,9 @@
using polynomial_function_diff [of f]
by (simp add: real_polynomial_function_eq)
-lemma real_polynomial_function_setsum [intro]:
- "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. setsum (f x) I)"
- using polynomial_function_setsum [of I f]
+lemma real_polynomial_function_sum [intro]:
+ "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. sum (f x) I)"
+ using polynomial_function_sum [of I f]
by (simp add: real_polynomial_function_eq)
lemma real_polynomial_function_power [intro]:
@@ -886,20 +886,20 @@
using g real_polynomial_function_compose [OF f]
by (auto simp: polynomial_function_def o_def)
-lemma setsum_max_0:
+lemma sum_max_0:
fixes x::real (*in fact "'a::comm_ring_1"*)
shows "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..m. x^i * a i)"
proof -
have "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..max m n. (if i \<le> m then x^i * a i else 0))"
- by (auto simp: algebra_simps intro: setsum.cong)
+ by (auto simp: algebra_simps intro: sum.cong)
also have "... = (\<Sum>i = 0..m. (if i \<le> m then x^i * a i else 0))"
- by (rule setsum.mono_neutral_right) auto
+ by (rule sum.mono_neutral_right) auto
also have "... = (\<Sum>i = 0..m. x^i * a i)"
- by (auto simp: algebra_simps intro: setsum.cong)
+ by (auto simp: algebra_simps intro: sum.cong)
finally show ?thesis .
qed
-lemma real_polynomial_function_imp_setsum:
+lemma real_polynomial_function_imp_sum:
assumes "real_polynomial_function f"
shows "\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i)"
using assms
@@ -925,8 +925,8 @@
then show ?case
apply (rule_tac x="\<lambda>i. (if i \<le> n1 then a1 i else 0) + (if i \<le> n2 then a2 i else 0)" in exI)
apply (rule_tac x="max n1 n2" in exI)
- using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1]
- apply (simp add: setsum.distrib algebra_simps max.commute)
+ using sum_max_0 [where m=n1 and n=n2] sum_max_0 [where m=n2 and n=n1]
+ apply (simp add: sum.distrib algebra_simps max.commute)
done
case (mult f1 f2)
then obtain a1 n1 a2 n2 where
@@ -944,11 +944,11 @@
done
qed
-lemma real_polynomial_function_iff_setsum:
+lemma real_polynomial_function_iff_sum:
"real_polynomial_function f \<longleftrightarrow> (\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i))"
apply (rule iffI)
- apply (erule real_polynomial_function_imp_setsum)
- apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum)
+ apply (erule real_polynomial_function_imp_sum)
+ apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_sum)
done
lemma polynomial_function_iff_Basis_inner:
@@ -969,7 +969,7 @@
apply (auto simp: real_polynomial_function_eq polynomial_function_mult)
done
then show "real_polynomial_function (h \<circ> f)"
- by (simp add: euclidean_representation_setsum_fun)
+ by (simp add: euclidean_representation_sum_fun)
qed
subsection \<open>Stone-Weierstrass theorem for polynomial functions\<close>
@@ -1058,8 +1058,8 @@
show ?thesis
apply (rule_tac p'="\<lambda>x. \<Sum>b\<in>Basis. qf b x" in that)
apply (force intro: qf)
- apply (subst euclidean_representation_setsum_fun [of p, symmetric])
- apply (auto intro: has_vector_derivative_setsum qf)
+ apply (subst euclidean_representation_sum_fun [of p, symmetric])
+ apply (auto intro: has_vector_derivative_sum qf)
done
qed
@@ -1068,14 +1068,14 @@
assumes "x \<noteq> y" shows "\<exists>f. real_polynomial_function f \<and> f x \<noteq> f y"
proof -
have "real_polynomial_function (\<lambda>u. \<Sum>b\<in>Basis. (inner (x-u) b)^2)"
- apply (rule real_polynomial_function_setsum)
+ apply (rule real_polynomial_function_sum)
apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff
const linear bounded_linear_inner_left)
done
then show ?thesis
apply (intro exI conjI, assumption)
using assms
- apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps)
+ apply (force simp add: euclidean_eq_iff [of x y] sum_nonneg_eq_0_iff algebra_simps)
done
qed
@@ -1117,14 +1117,14 @@
done
have "polynomial_function (\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b)"
using pf
- by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq)
+ by (simp add: polynomial_function_sum polynomial_function_mult real_polynomial_function_eq)
moreover
{ fix x
assume "x \<in> S"
have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) \<le> (\<Sum>b\<in>Basis. norm ((f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b))"
- by (rule norm_setsum)
+ by (rule norm_sum)
also have "... < of_nat DIM('b) * (e / DIM('b))"
- apply (rule setsum_bounded_above_strict)
+ apply (rule sum_bounded_above_strict)
apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \<open>x \<in> S\<close>)
apply (rule DIM_positive)
done
@@ -1134,9 +1134,9 @@
}
ultimately
show ?thesis
- apply (subst euclidean_representation_setsum_fun [of f, symmetric])
+ apply (subst euclidean_representation_sum_fun [of f, symmetric])
apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b" in exI)
- apply (auto simp: setsum_subtractf [symmetric])
+ apply (auto simp: sum_subtractf [symmetric])
done
qed
@@ -1247,7 +1247,7 @@
fixes i :: "'a::euclidean_space"
shows "polynomial_function g \<Longrightarrow> polynomial_function (\<lambda>x. g x \<bullet> i)"
apply (subst euclidean_representation [where x=i, symmetric])
- apply (force simp: inner_setsum_right polynomial_function_iff_Basis_inner polynomial_function_setsum)
+ apply (force simp: inner_sum_right polynomial_function_iff_Basis_inner polynomial_function_sum)
done
text\<open> Differentiability of real and vector polynomial functions.\<close>
@@ -1292,18 +1292,18 @@
shows "(\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = x"
proof (rule vector_eq_dot_span [OF _ \<open>x \<in> span B\<close>])
show "(\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) \<in> span B"
- by (simp add: span_clauses span_setsum)
+ by (simp add: span_clauses span_sum)
show "i \<bullet> (\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = i \<bullet> x" if "i \<in> B" for i
proof -
have [simp]: "i \<bullet> j = (if j = i then 1 else 0)" if "j \<in> B" for j
using B 1 that \<open>i \<in> B\<close>
by (force simp: norm_eq_1 orthogonal_def pairwise_def)
have "i \<bullet> (\<Sum>i\<in>B. (x \<bullet> i) *\<^sub>R i) = (\<Sum>j\<in>B. x \<bullet> j * (i \<bullet> j))"
- by (simp add: inner_setsum_right)
+ by (simp add: inner_sum_right)
also have "... = (\<Sum>j\<in>B. if j = i then x \<bullet> i else 0)"
- by (rule setsum.cong; simp)
+ by (rule sum.cong; simp)
also have "... = i \<bullet> x"
- by (simp add: \<open>finite B\<close> that inner_commute setsum.delta)
+ by (simp add: \<open>finite B\<close> that inner_commute sum.delta)
finally show ?thesis .
qed
qed
@@ -1343,11 +1343,11 @@
show ?thesis
proof
show "polynomial_function (\<lambda>x. \<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i)"
- apply (rule polynomial_function_setsum)
+ apply (rule polynomial_function_sum)
apply (simp add: \<open>finite B\<close>)
using \<open>polynomial_function g\<close> by auto
show "(\<lambda>x. \<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i) ` S \<subseteq> T"
- using \<open>B \<subseteq> T\<close> by (blast intro: subspace_setsum subspace_mul \<open>subspace T\<close>)
+ using \<open>B \<subseteq> T\<close> by (blast intro: subspace_sum subspace_mul \<open>subspace T\<close>)
show "norm (f x - (\<Sum>i\<in>B. (g x \<bullet> i) *\<^sub>R i)) < e" if "x \<in> S" for x
proof -
have orth': "pairwise (\<lambda>i j. orthogonal ((f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i)
@@ -1357,11 +1357,11 @@
done
then have "(norm (\<Sum>i\<in>B. (f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i))\<^sup>2 =
(\<Sum>i\<in>B. (norm ((f x \<bullet> i) *\<^sub>R i - (g x \<bullet> i) *\<^sub>R i))\<^sup>2)"
- by (simp add: norm_setsum_Pythagorean [OF \<open>finite B\<close> orth'])
+ by (simp add: norm_sum_Pythagorean [OF \<open>finite B\<close> orth'])
also have "... = (\<Sum>i\<in>B. (norm (((f x - g x) \<bullet> i) *\<^sub>R i))\<^sup>2)"
by (simp add: algebra_simps)
also have "... \<le> (\<Sum>i\<in>B. (norm (f x - g x))\<^sup>2)"
- apply (rule setsum_mono)
+ apply (rule sum_mono)
apply (simp add: B1)
apply (rule order_trans [OF Cauchy_Schwarz_ineq])
by (simp add: B1 dot_square_norm)
@@ -1380,7 +1380,7 @@
apply (rule power2_less_imp_less)
using \<open>0 < e\<close> by auto
then show ?thesis
- using fx that by (simp add: setsum_subtractf)
+ using fx that by (simp add: sum_subtractf)
qed
qed
qed
--- a/src/HOL/Analysis/ex/Approximations.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Analysis/ex/Approximations.thy Mon Oct 17 14:37:32 2016 +0200
@@ -19,7 +19,7 @@
by (simp, simp, simp_all only: numeral_eq_Suc fact_Suc,
simp only: numeral_eq_Suc [symmetric] of_nat_numeral)
-lemma setsum_poly_horner_expand:
+lemma sum_poly_horner_expand:
"(\<Sum>k<(numeral n::nat). f k * x^k) = f 0 + (\<Sum>k<pred_numeral n. f (k+1) * x^k) * x"
"(\<Sum>k<Suc 0. f k * x^k) = (f 0 :: 'a :: semiring_1)"
"(\<Sum>k<(0::nat). f k * x^k) = 0"
@@ -27,10 +27,10 @@
{
fix m :: nat
have "(\<Sum>k<Suc m. f k * x^k) = f 0 + (\<Sum>k=Suc 0..<Suc m. f k * x^k)"
- by (subst atLeast0LessThan [symmetric], subst setsum_head_upt_Suc) simp_all
+ by (subst atLeast0LessThan [symmetric], subst sum_head_upt_Suc) simp_all
also have "(\<Sum>k=Suc 0..<Suc m. f k * x^k) = (\<Sum>k<m. f (k+1) * x^k) * x"
- by (subst setsum_shift_bounds_Suc_ivl)
- (simp add: setsum_distrib_right algebra_simps atLeast0LessThan power_commutes)
+ by (subst sum_shift_bounds_Suc_ivl)
+ (simp add: sum_distrib_right algebra_simps atLeast0LessThan power_commutes)
finally have "(\<Sum>k<Suc m. f k * x ^ k) = f 0 + (\<Sum>k<m. f (k + 1) * x ^ k) * x" .
}
from this[of "pred_numeral n"]
@@ -181,9 +181,9 @@
proof -
have "real (\<Sum>k\<le>n. \<Prod>{k+1..n}) = (\<Sum>k\<le>n. \<Prod>i=k+1..n. real i)" by simp
also have "\<dots> / fact n = (\<Sum>k\<le>n. 1 / (fact n / (\<Prod>i=k+1..n. real i)))"
- by (simp add: setsum_divide_distrib)
+ by (simp add: sum_divide_distrib)
also have "\<dots> = (\<Sum>k\<le>n. 1 / fact k)"
- proof (intro setsum.cong refl)
+ proof (intro sum.cong refl)
fix k assume k: "k \<in> {..n}"
have "fact n = (\<Prod>i=1..n. real i)" by (simp add: fact_setprod)
also from k have "{1..n} = {1..k} \<union> {k+1..n}" by auto
@@ -199,7 +199,7 @@
lemma euler_approx_aux_Suc:
"euler_approx_aux (Suc m) = 1 + Suc m * euler_approx_aux m"
unfolding euler_approx_aux_def
- by (subst setsum_distrib_left) (simp add: atLeastAtMostSuc_conv)
+ by (subst sum_distrib_left) (simp add: atLeastAtMostSuc_conv)
lemma eval_euler_approx_aux:
"euler_approx_aux 0 = 1"
@@ -209,7 +209,7 @@
proof -
have A: "euler_approx_aux (Suc m) = 1 + Suc m * euler_approx_aux m" for m :: nat
unfolding euler_approx_aux_def
- by (subst setsum_distrib_left) (simp add: atLeastAtMostSuc_conv)
+ by (subst sum_distrib_left) (simp add: atLeastAtMostSuc_conv)
show ?th by (subst numeral_eq_Suc, subst A, subst numeral_eq_Suc [symmetric]) simp
qed (simp_all add: euler_approx_aux_def)
@@ -281,7 +281,7 @@
y_def [symmetric] d_def [symmetric])
also have "2 * y * (\<Sum>k<n. inverse (real (2 * k + 1)) * y\<^sup>2 ^ k) =
(\<Sum>k<n. 2 * y^(2*k+1) / (real (2 * k + 1)))"
- by (subst setsum_distrib_left, simp, subst power_mult)
+ by (subst sum_distrib_left, simp, subst power_mult)
(simp_all add: divide_simps mult_ac power_mult)
finally show ?case by (simp only: d_def y_def approx_def)
qed
@@ -297,7 +297,7 @@
apply (simp, simp add: ln_approx_aux1_def)
apply (simp add: ln_approx_aux2_def power2_eq_square power_divide)
apply (simp add: ln_approx_aux3_def power2_eq_square)
- apply (simp add: setsum_poly_horner_expand)
+ apply (simp add: sum_poly_horner_expand)
done
lemma ln_2_128:
@@ -380,7 +380,7 @@
from sums_split_initial_segment[OF this, of n]
have "(\<lambda>i. x * ((- x\<^sup>2) ^ (i + n) / real (2 * (i + n) + 1))) sums
(arctan x - arctan_approx n x)"
- by (simp add: arctan_approx_def setsum_distrib_left)
+ by (simp add: arctan_approx_def sum_distrib_left)
from sums_group[OF this, of 2] assms
have sums: "(\<lambda>k. x * (x\<^sup>2)^n * (x^4)^k * c k) sums (arctan x - arctan_approx n x)"
by (simp add: algebra_simps power_add power_mult [symmetric] c_def)
@@ -420,10 +420,10 @@
fix m :: nat
have "(\<Sum>k<Suc m. inverse (f k * x^k)) =
inverse (f 0) + (\<Sum>k=Suc 0..<Suc m. inverse (f k * x^k))"
- by (subst atLeast0LessThan [symmetric], subst setsum_head_upt_Suc) simp_all
+ by (subst atLeast0LessThan [symmetric], subst sum_head_upt_Suc) simp_all
also have "(\<Sum>k=Suc 0..<Suc m. inverse (f k * x^k)) = (\<Sum>k<m. inverse (f (k+1) * x^k)) / x"
- by (subst setsum_shift_bounds_Suc_ivl)
- (simp add: setsum_distrib_left divide_inverse algebra_simps
+ by (subst sum_shift_bounds_Suc_ivl)
+ (simp add: sum_distrib_left divide_inverse algebra_simps
atLeast0LessThan power_commutes)
finally have "(\<Sum>k<Suc m. inverse (f k) * inverse (x ^ k)) =
inverse (f 0) + (\<Sum>k<m. inverse (f (k + 1)) * inverse (x ^ k)) / x" by simp
--- a/src/HOL/Binomial.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Binomial.thy Mon Oct 17 14:37:32 2016 +0200
@@ -341,17 +341,17 @@
by (rule distrib_right)
also have "\<dots> = (\<Sum>k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) +
(\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n - k + 1))"
- by (auto simp add: setsum_distrib_left ac_simps)
+ by (auto simp add: sum_distrib_left ac_simps)
also have "\<dots> = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n + 1 - k)) +
(\<Sum>k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n + 1 - k))"
- by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps del: setsum_cl_ivl_Suc)
+ by (simp add:sum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps del: sum_cl_ivl_Suc)
also have "\<dots> = a^(n + 1) + b^(n + 1) +
(\<Sum>k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n + 1 - k)) +
(\<Sum>k=1..n. of_nat (n choose k) * a^k * b^(n + 1 - k))"
by (simp add: decomp2)
also have "\<dots> = a^(n + 1) + b^(n + 1) +
(\<Sum>k=1..n. of_nat (n + 1 choose k) * a^k * b^(n + 1 - k))"
- by (auto simp add: field_simps setsum.distrib [symmetric] choose_reduce_nat)
+ by (auto simp add: field_simps sum.distrib [symmetric] choose_reduce_nat)
also have "\<dots> = (\<Sum>k=0..n+1. of_nat (n + 1 choose k) * a^k * b^(n + 1 - k))"
using decomp by (simp add: field_simps)
finally show ?case
@@ -362,7 +362,7 @@
corollary binomial: "(a + b :: nat)^n = (\<Sum>k=0..n. (of_nat (n choose k)) * a^k * b^(n - k))"
using binomial_ring [of "int a" "int b" n]
by (simp only: of_nat_add [symmetric] of_nat_mult [symmetric] of_nat_power [symmetric]
- of_nat_setsum [symmetric] of_nat_eq_iff of_nat_id)
+ of_nat_sum [symmetric] of_nat_eq_iff of_nat_id)
lemma binomial_fact_lemma: "k \<le> n \<Longrightarrow> fact k * fact (n - k) * (n choose k) = fact n"
proof (induct n arbitrary: k rule: nat_less_induct)
@@ -459,11 +459,11 @@
proof -
have "2 ^ n = (\<Sum>i\<le>n. of_nat (n choose i)) + (\<Sum>i\<le>n. (-1) ^ i * of_nat (n choose i) :: 'a)"
using choose_row_sum[of n]
- by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_setsum[symmetric])
+ by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_sum[symmetric])
also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) + (-1) ^ i * of_nat (n choose i))"
- by (simp add: setsum.distrib)
+ by (simp add: sum.distrib)
also have "\<dots> = 2 * (\<Sum>i\<le>n. if even i then of_nat (n choose i) else 0)"
- by (subst setsum_distrib_left, intro setsum.cong) simp_all
+ by (subst sum_distrib_left, intro sum.cong) simp_all
finally show ?thesis ..
qed
@@ -473,11 +473,11 @@
proof -
have "2 ^ n = (\<Sum>i\<le>n. of_nat (n choose i)) - (\<Sum>i\<le>n. (-1) ^ i * of_nat (n choose i) :: 'a)"
using choose_row_sum[of n]
- by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_setsum[symmetric])
+ by (simp add: choose_alternating_sum assms atLeast0AtMost of_nat_sum[symmetric])
also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) - (-1) ^ i * of_nat (n choose i))"
- by (simp add: setsum_subtractf)
+ by (simp add: sum_subtractf)
also have "\<dots> = 2 * (\<Sum>i\<le>n. if odd i then of_nat (n choose i) else 0)"
- by (subst setsum_distrib_left, intro setsum.cong) simp_all
+ by (subst sum_distrib_left, intro sum.cong) simp_all
finally show ?thesis ..
qed
@@ -490,7 +490,7 @@
shows "(\<Sum>k=0..m. (n - k) choose (m - k)) = Suc n choose m"
proof -
have "(\<Sum>k=0..m. (n-k) choose (m - k)) = (\<Sum>k=0..m. (n - m + k) choose k)"
- using setsum.atLeast_atMost_rev [of "\<lambda>k. (n - k) choose (m - k)" 0 m] assms
+ using sum.atLeast_atMost_rev [of "\<lambda>k. (n - k) choose (m - k)" 0 m] assms
by simp
also have "\<dots> = Suc (n - m + m) choose m"
by (rule sum_choose_lower)
@@ -914,7 +914,7 @@
have "(\<Sum>i\<le>n. i * (n choose i)) = (\<Sum>i\<le>Suc m. i * (Suc m choose i))"
by (simp add: Suc)
also have "\<dots> = Suc m * 2 ^ m"
- by (simp only: setsum_atMost_Suc_shift Suc_times_binomial setsum_distrib_left[symmetric])
+ by (simp only: sum_atMost_Suc_shift Suc_times_binomial sum_distrib_left[symmetric])
(simp add: choose_row_sum')
finally show ?thesis
using Suc by simp
@@ -934,9 +934,9 @@
(\<Sum>i\<le>Suc m. (-1) ^ i * of_nat i * of_nat (Suc m choose i))"
by (simp add: Suc)
also have "\<dots> = (\<Sum>i\<le>m. (-1) ^ (Suc i) * of_nat (Suc i * (Suc m choose Suc i)))"
- by (simp only: setsum_atMost_Suc_shift setsum_distrib_left[symmetric] mult_ac of_nat_mult) simp
+ by (simp only: sum_atMost_Suc_shift sum_distrib_left[symmetric] mult_ac of_nat_mult) simp
also have "\<dots> = - of_nat (Suc m) * (\<Sum>i\<le>m. (-1) ^ i * of_nat (m choose i))"
- by (subst setsum_distrib_left, rule setsum.cong[OF refl], subst Suc_times_binomial)
+ by (subst sum_distrib_left, rule sum.cong[OF refl], subst Suc_times_binomial)
(simp add: algebra_simps)
also have "(\<Sum>i\<le>m. (-1 :: 'a) ^ i * of_nat ((m choose i))) = 0"
using choose_alternating_sum[OF \<open>m > 0\<close>] by simp
@@ -948,15 +948,15 @@
proof (induct n arbitrary: r)
case 0
have "(\<Sum>k\<le>r. (m choose k) * (0 choose (r - k))) = (\<Sum>k\<le>r. if k = r then (m choose k) else 0)"
- by (intro setsum.cong) simp_all
+ by (intro sum.cong) simp_all
also have "\<dots> = m choose r"
- by (simp add: setsum.delta)
+ by (simp add: sum.delta)
finally show ?case
by simp
next
case (Suc n r)
show ?case
- by (cases r) (simp_all add: Suc [symmetric] algebra_simps setsum.distrib Suc_diff_le)
+ by (cases r) (simp_all add: Suc [symmetric] algebra_simps sum.distrib Suc_diff_le)
qed
lemma choose_square_sum: "(\<Sum>k\<le>n. (n choose k)^2) = ((2*n) choose n)"
@@ -975,24 +975,24 @@
(\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
((\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
pochhammer b (Suc n))"
- by (subst setsum_atMost_Suc_shift) (simp add: ring_distribs setsum.distrib)
+ by (subst sum_atMost_Suc_shift) (simp add: ring_distribs sum.distrib)
also have "(\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a (Suc i) * pochhammer b (n - i)) =
a * pochhammer ((a + 1) + b) n"
- by (subst Suc) (simp add: setsum_distrib_left pochhammer_rec mult_ac)
+ by (subst Suc) (simp add: sum_distrib_left pochhammer_rec mult_ac)
also have "(\<Sum>i\<le>n. of_nat (n choose Suc i) * pochhammer a (Suc i) * pochhammer b (n - i)) +
pochhammer b (Suc n) =
(\<Sum>i=0..Suc n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
- apply (subst setsum_head_Suc)
+ apply (subst sum_head_Suc)
apply simp
- apply (subst setsum_shift_bounds_cl_Suc_ivl)
+ apply (subst sum_shift_bounds_cl_Suc_ivl)
apply (simp add: atLeast0AtMost)
done
also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc n - i))"
- using Suc by (intro setsum.mono_neutral_right) (auto simp: not_le binomial_eq_0)
+ using Suc by (intro sum.mono_neutral_right) (auto simp: not_le binomial_eq_0)
also have "\<dots> = (\<Sum>i\<le>n. of_nat (n choose i) * pochhammer a i * pochhammer b (Suc (n - i)))"
- by (intro setsum.cong) (simp_all add: Suc_diff_le)
+ by (intro sum.cong) (simp_all add: Suc_diff_le)
also have "\<dots> = b * pochhammer (a + (b + 1)) n"
- by (subst Suc) (simp add: setsum_distrib_left mult_ac pochhammer_rec)
+ by (subst Suc) (simp add: sum_distrib_left mult_ac pochhammer_rec)
also have "a * pochhammer ((a + 1) + b) n + b * pochhammer (a + (b + 1)) n =
pochhammer (a + b) (Suc n)"
by (simp add: pochhammer_rec algebra_simps)
@@ -1193,7 +1193,7 @@
(is "?lhs = ?rhs")
proof -
have "?lhs = (\<Sum>k\<le>m. -(r + 1) + of_nat k gchoose k)"
- by (intro setsum.cong[OF refl]) (subst gbinomial_negated_upper, simp add: power_mult_distrib)
+ by (intro sum.cong[OF refl]) (subst gbinomial_negated_upper, simp add: power_mult_distrib)
also have "\<dots> = - r + of_nat m gchoose m"
by (subst gbinomial_parallel_sum) simp
also have "\<dots> = ?rhs"
@@ -1218,7 +1218,7 @@
finally show ?case .
qed
-lemma setsum_bounds_lt_plus1: "(\<Sum>k<mm. f (Suc k)) = (\<Sum>k=1..mm. f k)"
+lemma sum_bounds_lt_plus1: "(\<Sum>k<mm. f (Suc k)) = (\<Sum>k=1..mm. f k)"
by (induct mm) simp_all
lemma gbinomial_partial_sum_poly:
@@ -1236,24 +1236,24 @@
unfolding S_def G_def ..
have "S (Suc mm) = G (Suc mm) 0 + (\<Sum>k=Suc 0..Suc mm. G (Suc mm) k)"
- using SG_def by (simp add: setsum_head_Suc atLeast0AtMost [symmetric])
+ using SG_def by (simp add: sum_head_Suc atLeast0AtMost [symmetric])
also have "(\<Sum>k=Suc 0..Suc mm. G (Suc mm) k) = (\<Sum>k=0..mm. G (Suc mm) (Suc k))"
- by (subst setsum_shift_bounds_cl_Suc_ivl) simp
+ by (subst sum_shift_bounds_cl_Suc_ivl) simp
also have "\<dots> = (\<Sum>k=0..mm. ((of_nat mm + r gchoose (Suc k)) +
(of_nat mm + r gchoose k)) * x^(Suc k) * y^(mm - k))"
unfolding G_def by (subst gbinomial_addition_formula) simp
also have "\<dots> = (\<Sum>k=0..mm. (of_nat mm + r gchoose (Suc k)) * x^(Suc k) * y^(mm - k)) +
(\<Sum>k=0..mm. (of_nat mm + r gchoose k) * x^(Suc k) * y^(mm - k))"
- by (subst setsum.distrib [symmetric]) (simp add: algebra_simps)
+ by (subst sum.distrib [symmetric]) (simp add: algebra_simps)
also have "(\<Sum>k=0..mm. (of_nat mm + r gchoose (Suc k)) * x^(Suc k) * y^(mm - k)) =
(\<Sum>k<Suc mm. (of_nat mm + r gchoose (Suc k)) * x^(Suc k) * y^(mm - k))"
by (simp only: atLeast0AtMost lessThan_Suc_atMost)
also have "\<dots> = (\<Sum>k<mm. (of_nat mm + r gchoose Suc k) * x^(Suc k) * y^(mm-k)) +
(of_nat mm + r gchoose (Suc mm)) * x^(Suc mm)"
(is "_ = ?A + ?B")
- by (subst setsum_lessThan_Suc) simp
+ by (subst sum_lessThan_Suc) simp
also have "?A = (\<Sum>k=1..mm. (of_nat mm + r gchoose k) * x^k * y^(mm - k + 1))"
- proof (subst setsum_bounds_lt_plus1 [symmetric], intro setsum.cong[OF refl], clarify)
+ proof (subst sum_bounds_lt_plus1 [symmetric], intro sum.cong[OF refl], clarify)
fix k
assume "k < mm"
then have "mm - k = mm - Suc k + 1"
@@ -1263,16 +1263,16 @@
by (simp only:)
qed
also have "\<dots> + ?B = y * (\<Sum>k=1..mm. (G mm k)) + (of_nat mm + r gchoose (Suc mm)) * x^(Suc mm)"
- unfolding G_def by (subst setsum_distrib_left) (simp add: algebra_simps)
+ unfolding G_def by (subst sum_distrib_left) (simp add: algebra_simps)
also have "(\<Sum>k=0..mm. (of_nat mm + r gchoose k) * x^(Suc k) * y^(mm - k)) = x * (S mm)"
- unfolding S_def by (subst setsum_distrib_left) (simp add: atLeast0AtMost algebra_simps)
+ unfolding S_def by (subst sum_distrib_left) (simp add: atLeast0AtMost algebra_simps)
also have "(G (Suc mm) 0) = y * (G mm 0)"
by (simp add: G_def)
finally have "S (Suc mm) =
y * (G mm 0 + (\<Sum>k=1..mm. (G mm k))) + (of_nat mm + r gchoose (Suc mm)) * x^(Suc mm) + x * (S mm)"
by (simp add: ring_distribs)
also have "G mm 0 + (\<Sum>k=1..mm. (G mm k)) = S mm"
- by (simp add: setsum_head_Suc[symmetric] SG_def atLeast0AtMost)
+ by (simp add: sum_head_Suc[symmetric] SG_def atLeast0AtMost)
finally have "S (Suc mm) = (x + y) * (S mm) + (of_nat mm + r gchoose (Suc mm)) * x^(Suc mm)"
by (simp add: algebra_simps)
also have "(of_nat mm + r gchoose (Suc mm)) = (-1) ^ (Suc mm) * (- r gchoose (Suc mm))"
@@ -1283,7 +1283,7 @@
also have "(x + y) * S mm + \<dots> = (x + y) * ?rhs mm + (- r gchoose (Suc mm)) * (- x)^Suc mm"
unfolding S_def by (subst Suc.IH) simp
also have "(x + y) * ?rhs mm = (\<Sum>n\<le>mm. ((- r gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n)))"
- by (subst setsum_distrib_left, rule setsum.cong) (simp_all add: Suc_diff_le)
+ by (subst sum_distrib_left, rule sum.cong) (simp_all add: Suc_diff_le)
also have "\<dots> + (-r gchoose (Suc mm)) * (-x)^Suc mm =
(\<Sum>n\<le>Suc mm. (- r gchoose n) * (- x) ^ n * (x + y) ^ (Suc mm - n))"
by simp
@@ -1296,7 +1296,7 @@
(\<Sum>k\<le>m. (of_nat k + r - 1 gchoose k) * x^k * (x + y)^(m-k))"
apply (subst gbinomial_partial_sum_poly)
apply (subst gbinomial_negated_upper)
- apply (intro setsum.cong, rule refl)
+ apply (intro sum.cong, rule refl)
apply (simp add: power_mult_distrib [symmetric])
done
@@ -1307,15 +1307,15 @@
also have "(\<Sum>k = 0..(2 * m + 1). (2 * m + 1 choose k)) =
(\<Sum>k = 0..m. (2 * m + 1 choose k)) +
(\<Sum>k = m+1..2*m+1. (2 * m + 1 choose k))"
- using setsum_ub_add_nat[of 0 m "\<lambda>k. 2 * m + 1 choose k" "m+1"]
+ using sum_ub_add_nat[of 0 m "\<lambda>k. 2 * m + 1 choose k" "m+1"]
by (simp add: mult_2)
also have "(\<Sum>k = m+1..2*m+1. (2 * m + 1 choose k)) =
(\<Sum>k = 0..m. (2 * m + 1 choose (k + (m + 1))))"
- by (subst setsum_shift_bounds_cl_nat_ivl [symmetric]) (simp add: mult_2)
+ by (subst sum_shift_bounds_cl_nat_ivl [symmetric]) (simp add: mult_2)
also have "\<dots> = (\<Sum>k = 0..m. (2 * m + 1 choose (m - k)))"
- by (intro setsum.cong[OF refl], subst binomial_symmetric) simp_all
+ by (intro sum.cong[OF refl], subst binomial_symmetric) simp_all
also have "\<dots> = (\<Sum>k = 0..m. (2 * m + 1 choose k))"
- using setsum.atLeast_atMost_rev [of "\<lambda>k. 2 * m + 1 choose (m - k)" 0 m]
+ using sum.atLeast_atMost_rev [of "\<lambda>k. 2 * m + 1 choose (m - k)" 0 m]
by simp
also have "\<dots> + \<dots> = 2 * \<dots>"
by simp
@@ -1345,7 +1345,7 @@
using gbinomial_partial_sum_poly_xpos[where x="1" and y="1" and r="of_nat m + 1" and m="m"]
by (simp add: add_ac)
also have "\<dots> = 2 ^ m * (\<Sum>k\<le>m. (of_nat (m + k) gchoose k) / 2 ^ k)"
- by (subst setsum_distrib_left) (simp add: algebra_simps power_diff)
+ by (subst sum_distrib_left) (simp add: algebra_simps power_diff)
finally show ?thesis
by (subst (asm) mult_left_cancel) simp_all
qed
@@ -1444,18 +1444,18 @@
by simp
also have "\<dots> = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (\<Sum>_\<in>\<Inter>I. (- 1) ^ (card I + 1)))"
(is "_ = nat ?rhs")
- by (subst setsum_distrib_left) simp
+ by (subst sum_distrib_left) simp
also have "?rhs = (\<Sum>(I, _)\<in>Sigma {I. I \<subseteq> A \<and> I \<noteq> {}} Inter. (- 1) ^ (card I + 1))"
- using assms by (subst setsum.Sigma) auto
+ using assms by (subst sum.Sigma) auto
also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:UNIV. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). (- 1) ^ (card I + 1))"
- by (rule setsum.reindex_cong [where l = "\<lambda>(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta)
+ by (rule sum.reindex_cong [where l = "\<lambda>(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta)
also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:\<Union>A. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). (- 1) ^ (card I + 1))"
using assms
- by (auto intro!: setsum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\<Union>A"])
+ by (auto intro!: sum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\<Union>A"])
also have "\<dots> = (\<Sum>x\<in>\<Union>A. (\<Sum>I|I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I. (- 1) ^ (card I + 1)))"
- using assms by (subst setsum.Sigma) auto
- also have "\<dots> = (\<Sum>_\<in>\<Union>A. 1)" (is "setsum ?lhs _ = _")
- proof (rule setsum.cong[OF refl])
+ using assms by (subst sum.Sigma) auto
+ also have "\<dots> = (\<Sum>_\<in>\<Union>A. 1)" (is "sum ?lhs _ = _")
+ proof (rule sum.cong[OF refl])
fix x
assume x: "x \<in> \<Union>A"
define K where "K = {X \<in> A. x \<in> X}"
@@ -1470,14 +1470,14 @@
simp add: card_gt_0_iff[folded Suc_le_eq]
dest: finite_subset intro: card_mono)
ultimately have "?lhs x = (\<Sum>(i, I)\<in>(SIGMA i:{1..card A}. ?I i). (- 1) ^ (i + 1))"
- by (rule setsum.reindex_cong [where l = snd]) fastforce
+ by (rule sum.reindex_cong [where l = snd]) fastforce
also have "\<dots> = (\<Sum>i=1..card A. (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. (- 1) ^ (i + 1)))"
- using assms by (subst setsum.Sigma) auto
+ using assms by (subst sum.Sigma) auto
also have "\<dots> = (\<Sum>i=1..card A. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1))"
- by (subst setsum_distrib_left) simp
+ by (subst sum_distrib_left) simp
also have "\<dots> = (\<Sum>i=1..card K. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> K \<and> card I = i. 1))"
(is "_ = ?rhs")
- proof (rule setsum.mono_neutral_cong_right[rule_format])
+ proof (rule sum.mono_neutral_cong_right[rule_format])
show "finite {1..card A}"
by simp
show "{1..card K} \<subseteq> {1..card A}"
@@ -1497,18 +1497,18 @@
fix i
have "(\<Sum>I | I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1) = (\<Sum>I | I \<subseteq> K \<and> card I = i. 1 :: int)"
(is "?lhs = ?rhs")
- by (rule setsum.cong) (auto simp add: K_def)
+ by (rule sum.cong) (auto simp add: K_def)
then show "(- 1) ^ (i + 1) * ?lhs = (- 1) ^ (i + 1) * ?rhs"
by simp
qed
also have "{I. I \<subseteq> K \<and> card I = 0} = {{}}"
using assms by (auto simp add: card_eq_0_iff K_def dest: finite_subset)
then have "?rhs = (\<Sum>i = 0..card K. (- 1) ^ (i + 1) * (\<Sum>I | I \<subseteq> K \<and> card I = i. 1 :: int)) + 1"
- by (subst (2) setsum_head_Suc) simp_all
+ by (subst (2) sum_head_Suc) simp_all
also have "\<dots> = (\<Sum>i = 0..card K. (- 1) * ((- 1) ^ i * int (card K choose i))) + 1"
using K by (subst n_subsets[symmetric]) simp_all
also have "\<dots> = - (\<Sum>i = 0..card K. (- 1) ^ i * int (card K choose i)) + 1"
- by (subst setsum_distrib_left[symmetric]) simp
+ by (subst sum_distrib_left[symmetric]) simp
also have "\<dots> = - ((-1 + 1) ^ card K) + 1"
by (subst binomial_ring) (simp add: ac_simps)
also have "\<dots> = 1"
--- a/src/HOL/Complex.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Complex.thy Mon Oct 17 14:37:32 2016 +0200
@@ -471,7 +471,7 @@
lemma complex_cnj_add [simp]: "cnj (x + y) = cnj x + cnj y"
by (simp add: complex_eq_iff)
-lemma cnj_setsum [simp]: "cnj (setsum f s) = (\<Sum>x\<in>s. cnj (f x))"
+lemma cnj_sum [simp]: "cnj (sum f s) = (\<Sum>x\<in>s. cnj (f x))"
by (induct s rule: infinite_finite_induct) auto
lemma complex_cnj_diff [simp]: "cnj (x - y) = cnj x - cnj y"
@@ -559,7 +559,7 @@
by (simp add: tendsto_iff dist_complex_def complex_cnj_diff [symmetric] del: complex_cnj_diff)
lemma sums_cnj: "((\<lambda>x. cnj(f x)) sums cnj l) \<longleftrightarrow> (f sums l)"
- by (simp add: sums_def lim_cnj cnj_setsum [symmetric] del: cnj_setsum)
+ by (simp add: sums_def lim_cnj cnj_sum [symmetric] del: cnj_sum)
subsection \<open>Basic Lemmas\<close>
@@ -630,14 +630,14 @@
lemma Im_divide_Reals: "r \<in> Reals \<Longrightarrow> Im (z / r) = Im z / Re r"
by (metis Im_divide_of_real of_real_Re)
-lemma Re_setsum[simp]: "Re (setsum f s) = (\<Sum>x\<in>s. Re (f x))"
+lemma Re_sum[simp]: "Re (sum f s) = (\<Sum>x\<in>s. Re (f x))"
by (induct s rule: infinite_finite_induct) auto
-lemma Im_setsum[simp]: "Im (setsum f s) = (\<Sum>x\<in>s. Im(f x))"
+lemma Im_sum[simp]: "Im (sum f s) = (\<Sum>x\<in>s. Im(f x))"
by (induct s rule: infinite_finite_induct) auto
lemma sums_complex_iff: "f sums x \<longleftrightarrow> ((\<lambda>x. Re (f x)) sums Re x) \<and> ((\<lambda>x. Im (f x)) sums Im x)"
- unfolding sums_def tendsto_complex_iff Im_setsum Re_setsum ..
+ unfolding sums_def tendsto_complex_iff Im_sum Re_sum ..
lemma summable_complex_iff: "summable f \<longleftrightarrow> summable (\<lambda>x. Re (f x)) \<and> summable (\<lambda>x. Im (f x))"
unfolding summable_def sums_complex_iff[abs_def] by (metis complex.sel)
@@ -1028,8 +1028,8 @@
and complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)"
and complex_eq_cancel_iff2: "(Complex x y = complex_of_real xa) = (x = xa \<and> y = 0)"
and complex_cn: "cnj (Complex a b) = Complex a (- b)"
- and Complex_setsum': "setsum (\<lambda>x. Complex (f x) 0) s = Complex (setsum f s) 0"
- and Complex_setsum: "Complex (setsum f s) 0 = setsum (\<lambda>x. Complex (f x) 0) s"
+ and Complex_sum': "sum (\<lambda>x. Complex (f x) 0) s = Complex (sum f s) 0"
+ and Complex_sum: "Complex (sum f s) 0 = sum (\<lambda>x. Complex (f x) 0) s"
and complex_of_real_def: "complex_of_real r = Complex r 0"
and complex_norm: "cmod (Complex x y) = sqrt (x\<^sup>2 + y\<^sup>2)"
by (simp_all add: norm_complex_def field_simps complex_eq_iff Re_divide Im_divide del: Complex_eq)
--- a/src/HOL/Datatype_Examples/Misc_Primrec.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Datatype_Examples/Misc_Primrec.thy Mon Oct 17 14:37:32 2016 +0200
@@ -45,7 +45,7 @@
primrec
hf_size :: "hfset \<Rightarrow> nat"
where
- "hf_size (HFset X) = 1 + setsum id (fset (fimage hf_size X))"
+ "hf_size (HFset X) = 1 + sum id (fset (fimage hf_size X))"
primrec rename_lam :: "(string \<Rightarrow> string) \<Rightarrow> lambda \<Rightarrow> lambda" where
"rename_lam f (Var s) = Var (f s)" |
--- a/src/HOL/Decision_Procs/Approximation.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Decision_Procs/Approximation.thy Mon Oct 17 14:37:32 2016 +0200
@@ -31,9 +31,9 @@
have shift_pow: "\<And>i. - (x * ((-1)^i * a (Suc i) * x ^ i)) = (-1)^(Suc i) * a (Suc i) * x ^ (Suc i)"
by auto
show ?thesis
- unfolding setsum_distrib_left shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
- setsum_head_upt_Suc[OF zero_less_Suc]
- setsum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n *a n * x^n"] by auto
+ unfolding sum_distrib_left shift_pow uminus_add_conv_diff [symmetric] sum_negf[symmetric]
+ sum_head_upt_Suc[OF zero_less_Suc]
+ sum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n *a n * x^n"] by auto
qed
lemma horner_schema:
@@ -514,7 +514,7 @@
proof -
have "(sqrt y * lb_arctan_horner prec n 1 y) \<le> ?S n"
using bounds(1) \<open>0 \<le> sqrt y\<close>
- apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
+ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric])
apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult)
apply (auto intro!: mult_left_mono)
done
@@ -527,7 +527,7 @@
have "arctan (sqrt y) \<le> ?S (Suc n)" using arctan_bounds ..
also have "\<dots> \<le> (sqrt y * ub_arctan_horner prec (Suc n) 1 y)"
using bounds(2)[of "Suc n"] \<open>0 \<le> sqrt y\<close>
- apply (simp only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
+ apply (simp only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric])
apply (simp only: mult.commute[where 'a=real] mult.commute[of _ "2::nat"] power_mult)
apply (auto intro!: mult_left_mono)
done
@@ -1116,7 +1116,7 @@
also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then (- 1) ^ (i div 2) / ((fact i)) * x ^ i else 0)"
unfolding sum_split_even_odd atLeast0LessThan ..
also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then (- 1) ^ (i div 2) / ((fact i)) else 0) * x ^ i)"
- by (rule setsum.cong) auto
+ by (rule sum.cong) auto
finally show ?thesis .
qed
@@ -1212,7 +1212,7 @@
from horner_bounds[where lb="lb_sin_cos_aux prec" and ub="ub_sin_cos_aux prec" and j'=0,
OF \<open>0 \<le> real_of_float (x * x)\<close> f_eq lb_sin_cos_aux.simps ub_sin_cos_aux.simps]
show "?lb" and "?ub" using \<open>0 \<le> real_of_float x\<close>
- apply (simp_all only: power_add power_one_right mult.assoc[symmetric] setsum_distrib_right[symmetric])
+ apply (simp_all only: power_add power_one_right mult.assoc[symmetric] sum_distrib_right[symmetric])
apply (simp_all only: mult.commute[where 'a=real] of_nat_fact)
apply (auto intro!: mult_left_mono simp add: power_mult power2_eq_square[of "real_of_float x"])
done
@@ -1230,7 +1230,7 @@
have "0 < x * x"
using \<open>0 < x\<close> by simp
- have setsum_morph: "(\<Sum>j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1)) =
+ have sum_morph: "(\<Sum>j = 0 ..< n. (- 1) ^ (((2 * j + 1) - Suc 0) div 2) / ((fact (2 * j + 1))) * x ^(2 * j + 1)) =
(\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * x ^ i)"
(is "?SUM = _") for x :: real and n
proof -
@@ -1241,7 +1241,7 @@
also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i)) * x ^ i)"
unfolding sum_split_even_odd atLeast0LessThan ..
also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then 0 else (- 1) ^ ((i - Suc 0) div 2) / ((fact i))) * x ^ i)"
- by (rule setsum.cong) auto
+ by (rule sum.cong) auto
finally show ?thesis .
qed
@@ -1273,7 +1273,7 @@
assume "even n"
have "(x * lb_sin_cos_aux prec n 2 1 (x * x)) \<le>
(\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)"
- using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding setsum_morph[symmetric] by auto
+ using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding sum_morph[symmetric] by auto
also have "\<dots> \<le> ?SUM" by auto
also have "\<dots> \<le> sin x"
proof -
@@ -1296,7 +1296,7 @@
also have "\<dots> \<le> (\<Sum> i = 0 ..< 2 * n. (if even(i) then 0 else ((- 1) ^ ((i - Suc 0) div 2))/((fact i))) * (real_of_float x) ^ i)"
by auto
also have "\<dots> \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))"
- using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding setsum_morph[symmetric] by auto
+ using sin_aux[OF \<open>0 \<le> real_of_float x\<close>] unfolding sum_morph[symmetric] by auto
finally have "sin x \<le> (x * ub_sin_cos_aux prec n 2 1 (x * x))" .
} note ub = this and lb
} note ub = this(1) and lb = this(2)
@@ -2192,8 +2192,8 @@
let "?s n" = "(- 1) ^ n * (1 / real (1 + n)) * (real_of_float x)^(Suc n)"
- have "?lb \<le> setsum ?s {0 ..< 2 * ev}"
- unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_distrib_right[symmetric]
+ have "?lb \<le> sum ?s {0 ..< 2 * ev}"
+ unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric]
unfolding mult.commute[of "real_of_float x"] ev
using horner_bounds(1)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x"
and lb="\<lambda>n i k x. lb_ln_horner prec n k x"
@@ -2205,10 +2205,10 @@
using ln_bounds(1)[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x < 1\<close>] by auto
finally show "?lb \<le> ?ln" .
- have "?ln \<le> setsum ?s {0 ..< 2 * od + 1}"
+ have "?ln \<le> sum ?s {0 ..< 2 * od + 1}"
using ln_bounds(2)[OF \<open>0 \<le> real_of_float x\<close> \<open>real_of_float x < 1\<close>] by auto
also have "\<dots> \<le> ?ub"
- unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq setsum_distrib_right[symmetric]
+ unfolding power_Suc2 mult.assoc[symmetric] times_float.rep_eq sum_distrib_right[symmetric]
unfolding mult.commute[of "real_of_float x"] od
using horner_bounds(2)[where G="\<lambda> i k. Suc k" and F="\<lambda>x. x" and f="\<lambda>x. x" and lb="\<lambda>n i k x. lb_ln_horner prec n k x" and ub="\<lambda>n i k x. ub_ln_horner prec n k x" and j'=1 and n="2*od+1",
OF \<open>0 \<le> real_of_float x\<close> refl lb_ln_horner.simps ub_ln_horner.simps] \<open>0 \<le> real_of_float x\<close>
@@ -3974,9 +3974,9 @@
have "\<And>k i. k < i \<Longrightarrow> {k ..< i} = insert k {Suc k ..< i}" by auto
hence setprod_head_Suc: "\<And>k i. \<Prod>{k ..< k + Suc i} = k * \<Prod>{Suc k ..< Suc k + i}"
by auto
- have setsum_move0: "\<And>k F. setsum F {0..<Suc k} = F 0 + setsum (\<lambda> k. F (Suc k)) {0..<k}"
- unfolding setsum_shift_bounds_Suc_ivl[symmetric]
- unfolding setsum_head_upt_Suc[OF zero_less_Suc] ..
+ have sum_move0: "\<And>k F. sum F {0..<Suc k} = F 0 + sum (\<lambda> k. F (Suc k)) {0..<k}"
+ unfolding sum_shift_bounds_Suc_ivl[symmetric]
+ unfolding sum_head_upt_Suc[OF zero_less_Suc] ..
define C where "C = xs!x - c"
{
@@ -3994,9 +3994,9 @@
also have "?X (Suc k) f n t * (xs!x - real_of_float c) * inverse (real k) + ?f 0 c =
(\<Sum> i = 0..<Suc n. inverse (real (\<Prod> j \<in> {k..<k+i}. j)) * ?f i c * (xs!x - c)^i) +
inverse (real (\<Prod> j \<in> {k..<k+Suc n}. j)) * ?f (Suc n) t * (xs!x - c)^Suc n" (is "_ = ?T")
- unfolding funpow_Suc C_def[symmetric] setsum_move0 setprod_head_Suc
+ unfolding funpow_Suc C_def[symmetric] sum_move0 setprod_head_Suc
by (auto simp add: algebra_simps)
- (simp only: mult.left_commute [of _ "inverse (real k)"] setsum_distrib_left [symmetric])
+ (simp only: mult.left_commute [of _ "inverse (real k)"] sum_distrib_left [symmetric])
finally have "?T \<in> {l .. u}" .
}
thus ?thesis using DERIV by blast
@@ -4046,7 +4046,7 @@
proof (cases "xs ! x = c")
case True
from True[symmetric] hyp[OF bnd_xs] Suc show ?thesis
- unfolding F_def Suc setsum_head_upt_Suc[OF zero_less_Suc] setsum_shift_bounds_Suc_ivl
+ unfolding F_def Suc sum_head_upt_Suc[OF zero_less_Suc] sum_shift_bounds_Suc_ivl
by auto
next
case False
--- a/src/HOL/Deriv.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Deriv.thy Mon Oct 17 14:37:32 2016 +0200
@@ -118,7 +118,7 @@
by (simp add: field_simps scaleR_add_right scaleR_diff_right)
qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear)
-lemma has_derivative_setsum[simp, derivative_intros]:
+lemma has_derivative_sum[simp, derivative_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i has_derivative f' i) F) \<Longrightarrow>
((\<lambda>x. \<Sum>i\<in>I. f i x) has_derivative (\<lambda>x. \<Sum>i\<in>I. f' i x)) F"
by (induct I rule: infinite_finite_induct) simp_all
@@ -368,7 +368,7 @@
using insert by (intro has_derivative_mult) auto
also have "?P = (\<lambda>y. \<Sum>i'\<in>insert i I. f' i' y * (\<Prod>j\<in>insert i I - {i'}. f j x))"
using insert(1,2)
- by (auto simp add: setsum_distrib_left insert_Diff_if intro!: ext setsum.cong)
+ by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong)
finally show ?case
using insert by simp
qed
@@ -678,10 +678,10 @@
((\<lambda>x. f x + g x) has_vector_derivative (f' + g')) net"
by (auto simp: has_vector_derivative_def scaleR_right_distrib)
-lemma has_vector_derivative_setsum[derivative_intros]:
+lemma has_vector_derivative_sum[derivative_intros]:
"(\<And>i. i \<in> I \<Longrightarrow> (f i has_vector_derivative f' i) net) \<Longrightarrow>
((\<lambda>x. \<Sum>i\<in>I. f i x) has_vector_derivative (\<Sum>i\<in>I. f' i)) net"
- by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_setsum_right intro!: derivative_eq_intros)
+ by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros)
lemma has_vector_derivative_diff[derivative_intros]:
"(f has_vector_derivative f') net \<Longrightarrow> (g has_vector_derivative g') net \<Longrightarrow>
@@ -841,11 +841,11 @@
lemma DERIV_unique: "DERIV f x :> D \<Longrightarrow> DERIV f x :> E \<Longrightarrow> D = E"
unfolding DERIV_def by (rule LIM_unique)
-lemma DERIV_setsum[derivative_intros]:
+lemma DERIV_sum[derivative_intros]:
"(\<And> n. n \<in> S \<Longrightarrow> ((\<lambda>x. f x n) has_field_derivative (f' x n)) F) \<Longrightarrow>
- ((\<lambda>x. setsum (f x) S) has_field_derivative setsum (f' x) S) F"
- by (rule has_derivative_imp_has_field_derivative [OF has_derivative_setsum])
- (auto simp: setsum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative)
+ ((\<lambda>x. sum (f x) S) has_field_derivative sum (f' x) S) F"
+ by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum])
+ (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative)
lemma DERIV_inverse'[derivative_intros]:
assumes "(f has_field_derivative D) (at x within s)"
--- a/src/HOL/Groups_Big.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Groups_Big.thy Mon Oct 17 14:37:32 2016 +0200
@@ -473,35 +473,35 @@
context comm_monoid_add
begin
-sublocale setsum: comm_monoid_set plus 0
- defines setsum = setsum.F ..
+sublocale sum: comm_monoid_set plus 0
+ defines sum = sum.F ..
-abbreviation Setsum ("\<Sum>_" [1000] 999)
- where "\<Sum>A \<equiv> setsum (\<lambda>x. x) A"
+abbreviation Sum ("\<Sum>_" [1000] 999)
+ where "\<Sum>A \<equiv> sum (\<lambda>x. x) A"
end
-text \<open>Now: lot's of fancy syntax. First, @{term "setsum (\<lambda>x. e) A"} is written \<open>\<Sum>x\<in>A. e\<close>.\<close>
+text \<open>Now: lot's of fancy syntax. First, @{term "sum (\<lambda>x. e) A"} is written \<open>\<Sum>x\<in>A. e\<close>.\<close>
syntax (ASCII)
- "_setsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add" ("(3SUM _:_./ _)" [0, 51, 10] 10)
+ "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add" ("(3SUM _:_./ _)" [0, 51, 10] 10)
syntax
- "_setsum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add" ("(2\<Sum>_\<in>_./ _)" [0, 51, 10] 10)
+ "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add" ("(2\<Sum>_\<in>_./ _)" [0, 51, 10] 10)
translations \<comment> \<open>Beware of argument permutation!\<close>
- "\<Sum>i\<in>A. b" \<rightleftharpoons> "CONST setsum (\<lambda>i. b) A"
+ "\<Sum>i\<in>A. b" \<rightleftharpoons> "CONST sum (\<lambda>i. b) A"
text \<open>Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter \<open>\<Sum>x|P. e\<close>.\<close>
syntax (ASCII)
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10)
+ "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0, 0, 10] 10)
syntax
- "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(2\<Sum>_ | (_)./ _)" [0, 0, 10] 10)
+ "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(2\<Sum>_ | (_)./ _)" [0, 0, 10] 10)
translations
- "\<Sum>x|P. t" => "CONST setsum (\<lambda>x. t) {x. P}"
+ "\<Sum>x|P. t" => "CONST sum (\<lambda>x. t) {x. P}"
print_translation \<open>
let
- fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
+ fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
if x <> y then raise Match
else
let
@@ -509,55 +509,55 @@
val t' = subst_bound (x', t);
val P' = subst_bound (x', P);
in
- Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
+ Syntax.const @{syntax_const "_qsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
end
- | setsum_tr' _ = raise Match;
-in [(@{const_syntax setsum}, K setsum_tr')] end
+ | sum_tr' _ = raise Match;
+in [(@{const_syntax sum}, K sum_tr')] end
\<close>
(* TODO generalization candidates *)
-lemma (in comm_monoid_add) setsum_image_gen:
+lemma (in comm_monoid_add) sum_image_gen:
assumes fin: "finite S"
- shows "setsum g S = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
+ shows "sum g S = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
proof -
have "{y. y\<in> f`S \<and> f x = y} = {f x}" if "x \<in> S" for x
using that by auto
- then have "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
+ then have "sum g S = sum (\<lambda>x. sum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
by simp
- also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
- by (rule setsum.commute_restrict [OF fin finite_imageI [OF fin]])
+ also have "\<dots> = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
+ by (rule sum.commute_restrict [OF fin finite_imageI [OF fin]])
finally show ?thesis .
qed
subsubsection \<open>Properties in more restricted classes of structures\<close>
-lemma setsum_Un:
- "finite A \<Longrightarrow> finite B \<Longrightarrow> setsum f (A \<union> B) = setsum f A + setsum f B - setsum f (A \<inter> B)"
+lemma sum_Un:
+ "finite A \<Longrightarrow> finite B \<Longrightarrow> sum f (A \<union> B) = sum f A + sum f B - sum f (A \<inter> B)"
for f :: "'b \<Rightarrow> 'a::ab_group_add"
- by (subst setsum.union_inter [symmetric]) (auto simp add: algebra_simps)
+ by (subst sum.union_inter [symmetric]) (auto simp add: algebra_simps)
-lemma setsum_Un2:
+lemma sum_Un2:
assumes "finite (A \<union> B)"
- shows "setsum f (A \<union> B) = setsum f (A - B) + setsum f (B - A) + setsum f (A \<inter> B)"
+ shows "sum f (A \<union> B) = sum f (A - B) + sum f (B - A) + sum f (A \<inter> B)"
proof -
have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
by auto
with assms show ?thesis
- by simp (subst setsum.union_disjoint, auto)+
+ by simp (subst sum.union_disjoint, auto)+
qed
-lemma setsum_diff1:
+lemma sum_diff1:
fixes f :: "'b \<Rightarrow> 'a::ab_group_add"
assumes "finite A"
- shows "setsum f (A - {a}) = (if a \<in> A then setsum f A - f a else setsum f A)"
+ shows "sum f (A - {a}) = (if a \<in> A then sum f A - f a else sum f A)"
using assms by induct (auto simp: insert_Diff_if)
-lemma setsum_diff:
+lemma sum_diff:
fixes f :: "'b \<Rightarrow> 'a::ab_group_add"
assumes "finite A" "B \<subseteq> A"
- shows "setsum f (A - B) = setsum f A - setsum f B"
+ shows "sum f (A - B) = sum f A - sum f B"
proof -
from assms(2,1) have "finite B" by (rule finite_subset)
from this \<open>B \<subseteq> A\<close>
@@ -568,18 +568,18 @@
next
case (insert x F)
with \<open>finite A\<close> \<open>finite B\<close> show ?case
- by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
+ by (simp add: Diff_insert[where a=x and B=F] sum_diff1 insert_absorb)
qed
qed
-lemma (in ordered_comm_monoid_add) setsum_mono:
+lemma (in ordered_comm_monoid_add) sum_mono:
"(\<And>i. i\<in>K \<Longrightarrow> f i \<le> g i) \<Longrightarrow> (\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
by (induct K rule: infinite_finite_induct) (use add_mono in auto)
-lemma (in strict_ordered_comm_monoid_add) setsum_strict_mono:
+lemma (in strict_ordered_comm_monoid_add) sum_strict_mono:
assumes "finite A" "A \<noteq> {}"
and "\<And>x. x \<in> A \<Longrightarrow> f x < g x"
- shows "setsum f A < setsum g A"
+ shows "sum f A < sum g A"
using assms
proof (induct rule: finite_ne_induct)
case singleton
@@ -589,31 +589,31 @@
then show ?case by (auto simp: add_strict_mono)
qed
-lemma setsum_strict_mono_ex1:
+lemma sum_strict_mono_ex1:
fixes f g :: "'i \<Rightarrow> 'a::ordered_cancel_comm_monoid_add"
assumes "finite A"
and "\<forall>x\<in>A. f x \<le> g x"
and "\<exists>a\<in>A. f a < g a"
- shows "setsum f A < setsum g A"
+ shows "sum f A < sum g A"
proof-
from assms(3) obtain a where a: "a \<in> A" "f a < g a" by blast
- have "setsum f A = setsum f ((A - {a}) \<union> {a})"
+ have "sum f A = sum f ((A - {a}) \<union> {a})"
by(simp add: insert_absorb[OF \<open>a \<in> A\<close>])
- also have "\<dots> = setsum f (A - {a}) + setsum f {a}"
- using \<open>finite A\<close> by(subst setsum.union_disjoint) auto
- also have "setsum f (A - {a}) \<le> setsum g (A - {a})"
- by (rule setsum_mono) (simp add: assms(2))
- also from a have "setsum f {a} < setsum g {a}" by simp
- also have "setsum g (A - {a}) + setsum g {a} = setsum g((A - {a}) \<union> {a})"
- using \<open>finite A\<close> by (subst setsum.union_disjoint[symmetric]) auto
- also have "\<dots> = setsum g A" by (simp add: insert_absorb[OF \<open>a \<in> A\<close>])
+ also have "\<dots> = sum f (A - {a}) + sum f {a}"
+ using \<open>finite A\<close> by(subst sum.union_disjoint) auto
+ also have "sum f (A - {a}) \<le> sum g (A - {a})"
+ by (rule sum_mono) (simp add: assms(2))
+ also from a have "sum f {a} < sum g {a}" by simp
+ also have "sum g (A - {a}) + sum g {a} = sum g((A - {a}) \<union> {a})"
+ using \<open>finite A\<close> by (subst sum.union_disjoint[symmetric]) auto
+ also have "\<dots> = sum g A" by (simp add: insert_absorb[OF \<open>a \<in> A\<close>])
finally show ?thesis
by (auto simp add: add_right_mono add_strict_left_mono)
qed
-lemma setsum_mono_inv:
+lemma sum_mono_inv:
fixes f g :: "'i \<Rightarrow> 'a :: ordered_cancel_comm_monoid_add"
- assumes eq: "setsum f I = setsum g I"
+ assumes eq: "sum f I = sum g I"
assumes le: "\<And>i. i \<in> I \<Longrightarrow> f i \<le> g i"
assumes i: "i \<in> I"
assumes I: "finite I"
@@ -622,46 +622,46 @@
assume "\<not> ?thesis"
with le[OF i] have "f i < g i" by simp
with i have "\<exists>i\<in>I. f i < g i" ..
- from setsum_strict_mono_ex1[OF I _ this] le have "setsum f I < setsum g I"
+ from sum_strict_mono_ex1[OF I _ this] le have "sum f I < sum g I"
by blast
with eq show False by simp
qed
-lemma member_le_setsum:
+lemma member_le_sum:
fixes f :: "_ \<Rightarrow> 'b::{semiring_1, ordered_comm_monoid_add}"
assumes le: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x"
and "i \<in> A"
and "finite A"
- shows "f i \<le> setsum f A"
+ shows "f i \<le> sum f A"
proof -
- have "f i \<le> setsum f (A \<inter> {i})"
+ have "f i \<le> sum f (A \<inter> {i})"
by (simp add: assms)
also have "... = (\<Sum>x\<in>A. if x \<in> {i} then f x else 0)"
- using assms setsum.inter_restrict by blast
- also have "... \<le> setsum f A"
- apply (rule setsum_mono)
+ using assms sum.inter_restrict by blast
+ also have "... \<le> sum f A"
+ apply (rule sum_mono)
apply (auto simp: le)
done
finally show ?thesis .
qed
-lemma setsum_negf: "(\<Sum>x\<in>A. - f x) = - (\<Sum>x\<in>A. f x)"
+lemma sum_negf: "(\<Sum>x\<in>A. - f x) = - (\<Sum>x\<in>A. f x)"
for f :: "'b \<Rightarrow> 'a::ab_group_add"
by (induct A rule: infinite_finite_induct) auto
-lemma setsum_subtractf: "(\<Sum>x\<in>A. f x - g x) = (\<Sum>x\<in>A. f x) - (\<Sum>x\<in>A. g x)"
+lemma sum_subtractf: "(\<Sum>x\<in>A. f x - g x) = (\<Sum>x\<in>A. f x) - (\<Sum>x\<in>A. g x)"
for f g :: "'b \<Rightarrow>'a::ab_group_add"
- using setsum.distrib [of f "- g" A] by (simp add: setsum_negf)
+ using sum.distrib [of f "- g" A] by (simp add: sum_negf)
-lemma setsum_subtractf_nat:
+lemma sum_subtractf_nat:
"(\<And>x. x \<in> A \<Longrightarrow> g x \<le> f x) \<Longrightarrow> (\<Sum>x\<in>A. f x - g x) = (\<Sum>x\<in>A. f x) - (\<Sum>x\<in>A. g x)"
for f g :: "'a \<Rightarrow> nat"
- by (induct A rule: infinite_finite_induct) (auto simp: setsum_mono)
+ by (induct A rule: infinite_finite_induct) (auto simp: sum_mono)
context ordered_comm_monoid_add
begin
-lemma setsum_nonneg: "\<forall>x\<in>A. 0 \<le> f x \<Longrightarrow> 0 \<le> setsum f A"
+lemma sum_nonneg: "\<forall>x\<in>A. 0 \<le> f x \<Longrightarrow> 0 \<le> sum f A"
proof (induct A rule: infinite_finite_induct)
case infinite
then show ?case by simp
@@ -670,11 +670,11 @@
then show ?case by simp
next
case (insert x F)
- then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
+ then have "0 + 0 \<le> f x + sum f F" by (blast intro: add_mono)
with insert show ?case by simp
qed
-lemma setsum_nonpos: "\<forall>x\<in>A. f x \<le> 0 \<Longrightarrow> setsum f A \<le> 0"
+lemma sum_nonpos: "\<forall>x\<in>A. f x \<le> 0 \<Longrightarrow> sum f A \<le> 0"
proof (induct A rule: infinite_finite_induct)
case infinite
then show ?case by simp
@@ -683,74 +683,74 @@
then show ?case by simp
next
case (insert x F)
- then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
+ then have "f x + sum f F \<le> 0 + 0" by (blast intro: add_mono)
with insert show ?case by simp
qed
-lemma setsum_nonneg_eq_0_iff:
- "finite A \<Longrightarrow> \<forall>x\<in>A. 0 \<le> f x \<Longrightarrow> setsum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
- by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff setsum_nonneg)
+lemma sum_nonneg_eq_0_iff:
+ "finite A \<Longrightarrow> \<forall>x\<in>A. 0 \<le> f x \<Longrightarrow> sum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"
+ by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff sum_nonneg)
-lemma setsum_nonneg_0:
+lemma sum_nonneg_0:
"finite s \<Longrightarrow> (\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0) \<Longrightarrow> (\<Sum> i \<in> s. f i) = 0 \<Longrightarrow> i \<in> s \<Longrightarrow> f i = 0"
- by (simp add: setsum_nonneg_eq_0_iff)
+ by (simp add: sum_nonneg_eq_0_iff)
-lemma setsum_nonneg_leq_bound:
+lemma sum_nonneg_leq_bound:
assumes "finite s" "\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0" "(\<Sum>i \<in> s. f i) = B" "i \<in> s"
shows "f i \<le> B"
proof -
from assms have "f i \<le> f i + (\<Sum>i \<in> s - {i}. f i)"
- by (intro add_increasing2 setsum_nonneg) auto
+ by (intro add_increasing2 sum_nonneg) auto
also have "\<dots> = B"
- using setsum.remove[of s i f] assms by simp
+ using sum.remove[of s i f] assms by simp
finally show ?thesis by auto
qed
-lemma setsum_mono2:
+lemma sum_mono2:
assumes fin: "finite B"
and sub: "A \<subseteq> B"
and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
- shows "setsum f A \<le> setsum f B"
+ shows "sum f A \<le> sum f B"
proof -
- have "setsum f A \<le> setsum f A + setsum f (B-A)"
- by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
- also from fin finite_subset[OF sub fin] have "\<dots> = setsum f (A \<union> (B-A))"
- by (simp add: setsum.union_disjoint del: Un_Diff_cancel)
+ have "sum f A \<le> sum f A + sum f (B-A)"
+ by(simp add: add_increasing2[OF sum_nonneg] nn Ball_def)
+ also from fin finite_subset[OF sub fin] have "\<dots> = sum f (A \<union> (B-A))"
+ by (simp add: sum.union_disjoint del: Un_Diff_cancel)
also from sub have "A \<union> (B-A) = B" by blast
finally show ?thesis .
qed
-lemma setsum_le_included:
+lemma sum_le_included:
assumes "finite s" "finite t"
and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"
- shows "setsum f s \<le> setsum g t"
+ shows "sum f s \<le> sum g t"
proof -
- have "setsum f s \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) s"
- proof (rule setsum_mono)
+ have "sum f s \<le> sum (\<lambda>y. sum g {x. x\<in>t \<and> i x = y}) s"
+ proof (rule sum_mono)
fix y
assume "y \<in> s"
with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto
- with assms show "f y \<le> setsum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
- using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro]
- by (auto intro!: setsum_mono2)
+ with assms show "f y \<le> sum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
+ using order_trans[of "?A (i z)" "sum g {z}" "?B (i z)", intro]
+ by (auto intro!: sum_mono2)
qed
- also have "\<dots> \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) (i ` t)"
- using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg)
- also have "\<dots> \<le> setsum g t"
- using assms by (auto simp: setsum_image_gen[symmetric])
+ also have "\<dots> \<le> sum (\<lambda>y. sum g {x. x\<in>t \<and> i x = y}) (i ` t)"
+ using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg)
+ also have "\<dots> \<le> sum g t"
+ using assms by (auto simp: sum_image_gen[symmetric])
finally show ?thesis .
qed
-lemma setsum_mono3: "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> \<forall>x\<in>B - A. 0 \<le> f x \<Longrightarrow> setsum f A \<le> setsum f B"
- by (rule setsum_mono2) auto
+lemma sum_mono3: "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> \<forall>x\<in>B - A. 0 \<le> f x \<Longrightarrow> sum f A \<le> sum f B"
+ by (rule sum_mono2) auto
end
-lemma (in canonically_ordered_monoid_add) setsum_eq_0_iff [simp]:
- "finite F \<Longrightarrow> (setsum f F = 0) = (\<forall>a\<in>F. f a = 0)"
- by (intro ballI setsum_nonneg_eq_0_iff zero_le)
+lemma (in canonically_ordered_monoid_add) sum_eq_0_iff [simp]:
+ "finite F \<Longrightarrow> (sum f F = 0) = (\<forall>a\<in>F. f a = 0)"
+ by (intro ballI sum_nonneg_eq_0_iff zero_le)
-lemma setsum_distrib_left: "r * setsum f A = setsum (\<lambda>n. r * f n) A"
+lemma sum_distrib_left: "r * sum f A = sum (\<lambda>n. r * f n) A"
for f :: "'a \<Rightarrow> 'b::semiring_0"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -763,7 +763,7 @@
then show ?case by (simp add: distrib_left)
qed
-lemma setsum_distrib_right: "setsum f A * r = (\<Sum>n\<in>A. f n * r)"
+lemma sum_distrib_right: "sum f A * r = (\<Sum>n\<in>A. f n * r)"
for r :: "'a::semiring_0"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -776,7 +776,7 @@
then show ?case by (simp add: distrib_right)
qed
-lemma setsum_divide_distrib: "setsum f A / r = (\<Sum>n\<in>A. f n / r)"
+lemma sum_divide_distrib: "sum f A / r = (\<Sum>n\<in>A. f n / r)"
for r :: "'a::field"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -789,7 +789,7 @@
then show ?case by (simp add: add_divide_distrib)
qed
-lemma setsum_abs[iff]: "\<bar>setsum f A\<bar> \<le> setsum (\<lambda>i. \<bar>f i\<bar>) A"
+lemma sum_abs[iff]: "\<bar>sum f A\<bar> \<le> sum (\<lambda>i. \<bar>f i\<bar>) A"
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -802,11 +802,11 @@
then show ?case by (auto intro: abs_triangle_ineq order_trans)
qed
-lemma setsum_abs_ge_zero[iff]: "0 \<le> setsum (\<lambda>i. \<bar>f i\<bar>) A"
+lemma sum_abs_ge_zero[iff]: "0 \<le> sum (\<lambda>i. \<bar>f i\<bar>) A"
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"
- by (simp add: setsum_nonneg)
+ by (simp add: sum_nonneg)
-lemma abs_setsum_abs[simp]: "\<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar> = (\<Sum>a\<in>A. \<bar>f a\<bar>)"
+lemma abs_sum_abs[simp]: "\<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar> = (\<Sum>a\<in>A. \<bar>f a\<bar>)"
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -823,39 +823,39 @@
finally show ?case .
qed
-lemma setsum_diff1_ring:
+lemma sum_diff1_ring:
fixes f :: "'b \<Rightarrow> 'a::ring"
assumes "finite A" "a \<in> A"
- shows "setsum f (A - {a}) = setsum f A - (f a)"
- unfolding setsum.remove [OF assms] by auto
+ shows "sum f (A - {a}) = sum f A - (f a)"
+ unfolding sum.remove [OF assms] by auto
-lemma setsum_product:
+lemma sum_product:
fixes f :: "'a \<Rightarrow> 'b::semiring_0"
- shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
- by (simp add: setsum_distrib_left setsum_distrib_right) (rule setsum.commute)
+ shows "sum f A * sum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
+ by (simp add: sum_distrib_left sum_distrib_right) (rule sum.commute)
-lemma setsum_mult_setsum_if_inj:
+lemma sum_mult_sum_if_inj:
fixes f :: "'a \<Rightarrow> 'b::semiring_0"
shows "inj_on (\<lambda>(a, b). f a * g b) (A \<times> B) \<Longrightarrow>
- setsum f A * setsum g B = setsum id {f a * g b |a b. a \<in> A \<and> b \<in> B}"
- by(auto simp: setsum_product setsum.cartesian_product intro!: setsum.reindex_cong[symmetric])
+ sum f A * sum g B = sum id {f a * g b |a b. a \<in> A \<and> b \<in> B}"
+ by(auto simp: sum_product sum.cartesian_product intro!: sum.reindex_cong[symmetric])
-lemma setsum_SucD: "setsum f A = Suc n \<Longrightarrow> \<exists>a\<in>A. 0 < f a"
+lemma sum_SucD: "sum f A = Suc n \<Longrightarrow> \<exists>a\<in>A. 0 < f a"
by (induct A rule: infinite_finite_induct) auto
-lemma setsum_eq_Suc0_iff:
- "finite A \<Longrightarrow> setsum f A = Suc 0 \<longleftrightarrow> (\<exists>a\<in>A. f a = Suc 0 \<and> (\<forall>b\<in>A. a \<noteq> b \<longrightarrow> f b = 0))"
+lemma sum_eq_Suc0_iff:
+ "finite A \<Longrightarrow> sum f A = Suc 0 \<longleftrightarrow> (\<exists>a\<in>A. f a = Suc 0 \<and> (\<forall>b\<in>A. a \<noteq> b \<longrightarrow> f b = 0))"
by (induct A rule: finite_induct) (auto simp add: add_is_1)
-lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
+lemmas sum_eq_1_iff = sum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
-lemma setsum_Un_nat:
- "finite A \<Longrightarrow> finite B \<Longrightarrow> setsum f (A \<union> B) = setsum f A + setsum f B - setsum f (A \<inter> B)"
+lemma sum_Un_nat:
+ "finite A \<Longrightarrow> finite B \<Longrightarrow> sum f (A \<union> B) = sum f A + sum f B - sum f (A \<inter> B)"
for f :: "'a \<Rightarrow> nat"
\<comment> \<open>For the natural numbers, we have subtraction.\<close>
- by (subst setsum.union_inter [symmetric]) (auto simp: algebra_simps)
+ by (subst sum.union_inter [symmetric]) (auto simp: algebra_simps)
-lemma setsum_diff1_nat: "setsum f (A - {a}) = (if a \<in> A then setsum f A - f a else setsum f A)"
+lemma sum_diff1_nat: "sum f (A - {a}) = (if a \<in> A then sum f A - f a else sum f A)"
for f :: "'a \<Rightarrow> nat"
proof (induct A rule: infinite_finite_induct)
case infinite
@@ -872,60 +872,60 @@
done
qed
-lemma setsum_diff_nat:
+lemma sum_diff_nat:
fixes f :: "'a \<Rightarrow> nat"
assumes "finite B" and "B \<subseteq> A"
- shows "setsum f (A - B) = setsum f A - setsum f B"
+ shows "sum f (A - B) = sum f A - sum f B"
using assms
proof induct
case empty
then show ?case by simp
next
case (insert x F)
- note IH = \<open>F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F\<close>
+ note IH = \<open>F \<subseteq> A \<Longrightarrow> sum f (A - F) = sum f A - sum f F\<close>
from \<open>x \<notin> F\<close> \<open>insert x F \<subseteq> A\<close> have "x \<in> A - F" by simp
- then have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
- by (simp add: setsum_diff1_nat)
+ then have A: "sum f ((A - F) - {x}) = sum f (A - F) - f x"
+ by (simp add: sum_diff1_nat)
from \<open>insert x F \<subseteq> A\<close> have "F \<subseteq> A" by simp
- with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
- with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
+ with IH have "sum f (A - F) = sum f A - sum f F" by simp
+ with A have B: "sum f ((A - F) - {x}) = sum f A - sum f F - f x"
by simp
from \<open>x \<notin> F\<close> have "A - insert x F = (A - F) - {x}" by auto
- with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
+ with B have C: "sum f (A - insert x F) = sum f A - sum f F - f x"
by simp
- from \<open>finite F\<close> \<open>x \<notin> F\<close> have "setsum f (insert x F) = setsum f F + f x"
+ from \<open>finite F\<close> \<open>x \<notin> F\<close> have "sum f (insert x F) = sum f F + f x"
by simp
- with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
+ with C have "sum f (A - insert x F) = sum f A - sum f (insert x F)"
by simp
then show ?case by simp
qed
-lemma setsum_comp_morphism:
- "h 0 = 0 \<Longrightarrow> (\<And>x y. h (x + y) = h x + h y) \<Longrightarrow> setsum (h \<circ> g) A = h (setsum g A)"
+lemma sum_comp_morphism:
+ "h 0 = 0 \<Longrightarrow> (\<And>x y. h (x + y) = h x + h y) \<Longrightarrow> sum (h \<circ> g) A = h (sum g A)"
by (induct A rule: infinite_finite_induct) simp_all
-lemma (in comm_semiring_1) dvd_setsum: "(\<And>a. a \<in> A \<Longrightarrow> d dvd f a) \<Longrightarrow> d dvd setsum f A"
+lemma (in comm_semiring_1) dvd_sum: "(\<And>a. a \<in> A \<Longrightarrow> d dvd f a) \<Longrightarrow> d dvd sum f A"
by (induct A rule: infinite_finite_induct) simp_all
-lemma (in ordered_comm_monoid_add) setsum_pos:
- "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> 0 < f i) \<Longrightarrow> 0 < setsum f I"
+lemma (in ordered_comm_monoid_add) sum_pos:
+ "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> 0 < f i) \<Longrightarrow> 0 < sum f I"
by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos)
-lemma (in ordered_comm_monoid_add) setsum_pos2:
+lemma (in ordered_comm_monoid_add) sum_pos2:
assumes I: "finite I" "i \<in> I" "0 < f i" "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"
- shows "0 < setsum f I"
+ shows "0 < sum f I"
proof -
- have "0 < f i + setsum f (I - {i})"
- using assms by (intro add_pos_nonneg setsum_nonneg) auto
- also have "\<dots> = setsum f I"
- using assms by (simp add: setsum.remove)
+ have "0 < f i + sum f (I - {i})"
+ using assms by (intro add_pos_nonneg sum_nonneg) auto
+ also have "\<dots> = sum f I"
+ using assms by (simp add: sum.remove)
finally show ?thesis .
qed
-lemma setsum_cong_Suc:
+lemma sum_cong_Suc:
assumes "0 \<notin> A" "\<And>x. Suc x \<in> A \<Longrightarrow> f (Suc x) = g (Suc x)"
- shows "setsum f A = setsum g A"
-proof (rule setsum.cong)
+ shows "sum f A = sum g A"
+proof (rule sum.cong)
fix x
assume "x \<in> A"
with assms(1) show "f x = g x"
@@ -933,9 +933,9 @@
qed simp_all
-subsubsection \<open>Cardinality as special case of @{const setsum}\<close>
+subsubsection \<open>Cardinality as special case of @{const sum}\<close>
-lemma card_eq_setsum: "card A = setsum (\<lambda>x. 1) A"
+lemma card_eq_sum: "card A = sum (\<lambda>x. 1) A"
proof -
have "plus \<circ> (\<lambda>_. Suc 0) = (\<lambda>_. Suc)"
by (simp add: fun_eq_iff)
@@ -944,43 +944,43 @@
then have "Finite_Set.fold (plus \<circ> (\<lambda>_. Suc 0)) 0 A = Finite_Set.fold (\<lambda>_. Suc) 0 A"
by (blast intro: fun_cong)
then show ?thesis
- by (simp add: card.eq_fold setsum.eq_fold)
+ by (simp add: card.eq_fold sum.eq_fold)
qed
-lemma setsum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat (card A) * y"
+lemma sum_constant [simp]: "(\<Sum>x \<in> A. y) = of_nat (card A) * y"
by (induct A rule: infinite_finite_induct) (auto simp: algebra_simps)
-lemma setsum_Suc: "setsum (\<lambda>x. Suc(f x)) A = setsum f A + card A"
- using setsum.distrib[of f "\<lambda>_. 1" A] by simp
+lemma sum_Suc: "sum (\<lambda>x. Suc(f x)) A = sum f A + card A"
+ using sum.distrib[of f "\<lambda>_. 1" A] by simp
-lemma setsum_bounded_above:
+lemma sum_bounded_above:
fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}"
assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> K"
- shows "setsum f A \<le> of_nat (card A) * K"
+ shows "sum f A \<le> of_nat (card A) * K"
proof (cases "finite A")
case True
then show ?thesis
- using le setsum_mono[where K=A and g = "\<lambda>x. K"] by simp
+ using le sum_mono[where K=A and g = "\<lambda>x. K"] by simp
next
case False
then show ?thesis by simp
qed
-lemma setsum_bounded_above_strict:
+lemma sum_bounded_above_strict:
fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}"
assumes "\<And>i. i\<in>A \<Longrightarrow> f i < K" "card A > 0"
- shows "setsum f A < of_nat (card A) * K"
- using assms setsum_strict_mono[where A=A and g = "\<lambda>x. K"]
+ shows "sum f A < of_nat (card A) * K"
+ using assms sum_strict_mono[where A=A and g = "\<lambda>x. K"]
by (simp add: card_gt_0_iff)
-lemma setsum_bounded_below:
+lemma sum_bounded_below:
fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}"
assumes le: "\<And>i. i\<in>A \<Longrightarrow> K \<le> f i"
- shows "of_nat (card A) * K \<le> setsum f A"
+ shows "of_nat (card A) * K \<le> sum f A"
proof (cases "finite A")
case True
then show ?thesis
- using le setsum_mono[where K=A and f = "\<lambda>x. K"] by simp
+ using le sum_mono[where K=A and f = "\<lambda>x. K"] by simp
next
case False
then show ?thesis by simp
@@ -994,33 +994,33 @@
have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)"
by simp
with assms show ?thesis
- by (simp add: card_eq_setsum setsum.UNION_disjoint del: setsum_constant)
+ by (simp add: card_eq_sum sum.UNION_disjoint del: sum_constant)
qed
lemma card_Union_disjoint:
"finite C \<Longrightarrow> \<forall>A\<in>C. finite A \<Longrightarrow> \<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {} \<Longrightarrow>
- card (\<Union>C) = setsum card C"
+ card (\<Union>C) = sum card C"
by (frule card_UN_disjoint [of C id]) simp_all
-lemma setsum_multicount_gen:
+lemma sum_multicount_gen:
assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
- shows "setsum (\<lambda>i. (card {j\<in>t. R i j})) s = setsum k t"
+ shows "sum (\<lambda>i. (card {j\<in>t. R i j})) s = sum k t"
(is "?l = ?r")
proof-
- have "?l = setsum (\<lambda>i. setsum (\<lambda>x.1) {j\<in>t. R i j}) s"
+ have "?l = sum (\<lambda>i. sum (\<lambda>x.1) {j\<in>t. R i j}) s"
by auto
also have "\<dots> = ?r"
- unfolding setsum.commute_restrict [OF assms(1-2)]
+ unfolding sum.commute_restrict [OF assms(1-2)]
using assms(3) by auto
finally show ?thesis .
qed
-lemma setsum_multicount:
+lemma sum_multicount:
assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
- shows "setsum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
+ shows "sum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
proof-
- have "?l = setsum (\<lambda>i. k) T"
- by (rule setsum_multicount_gen) (auto simp: assms)
+ have "?l = sum (\<lambda>i. k) T"
+ by (rule sum_multicount_gen) (auto simp: assms)
also have "\<dots> = ?r" by (simp add: mult.commute)
finally show ?thesis by auto
qed
@@ -1030,7 +1030,7 @@
lemma card_SigmaI [simp]:
"finite A \<Longrightarrow> \<forall>a\<in>A. finite (B a) \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
- by (simp add: card_eq_setsum setsum.Sigma del: setsum_constant)
+ by (simp add: card_eq_sum sum.Sigma del: sum_constant)
(*
lemma SigmaI_insert: "y \<notin> A ==>
@@ -1187,14 +1187,14 @@
qed
qed
-lemma setsum_zero_power [simp]: "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"
+lemma sum_zero_power [simp]: "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"
for c :: "nat \<Rightarrow> 'a::division_ring"
by (induct A rule: infinite_finite_induct) auto
-lemma setsum_zero_power' [simp]:
+lemma sum_zero_power' [simp]:
"(\<Sum>i\<in>A. c i * 0^i / d i) = (if finite A \<and> 0 \<in> A then c 0 / d 0 else 0)"
for c :: "nat \<Rightarrow> 'a::field"
- using setsum_zero_power [of "\<lambda>i. c i / d i" A] by auto
+ using sum_zero_power [of "\<lambda>i. c i / d i" A] by auto
lemma (in field) setprod_inversef:
"finite A \<Longrightarrow> setprod (inverse \<circ> f) A = inverse (setprod f A)"
@@ -1256,7 +1256,7 @@
for f :: "'a \<Rightarrow> 'b::comm_semiring_1"
by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib)
-lemma power_setsum: "c ^ (\<Sum>a\<in>A. f a) = (\<Prod>a\<in>A. c ^ f a)"
+lemma power_sum: "c ^ (\<Sum>a\<in>A. f a) = (\<Prod>a\<in>A. c ^ f a)"
by (induct A rule: infinite_finite_induct) (simp_all add: power_add)
lemma setprod_gen_delta:
@@ -1291,20 +1291,20 @@
qed
qed
-lemma setsum_image_le:
+lemma sum_image_le:
fixes g :: "'a \<Rightarrow> 'b::ordered_ab_group_add"
assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> g(f i)"
- shows "setsum g (f ` I) \<le> setsum (g \<circ> f) I"
+ shows "sum g (f ` I) \<le> sum (g \<circ> f) I"
using assms
proof induction
case empty
then show ?case by auto
next
case (insert x F) then
- have "setsum g (f ` insert x F) = setsum g (insert (f x) (f ` F))" by simp
- also have "\<dots> \<le> g (f x) + setsum g (f ` F)"
- by (simp add: insert setsum.insert_if)
- also have "\<dots> \<le> setsum (g \<circ> f) (insert x F)"
+ have "sum g (f ` insert x F) = sum g (insert (f x) (f ` F))" by simp
+ also have "\<dots> \<le> g (f x) + sum g (f ` F)"
+ by (simp add: insert sum.insert_if)
+ also have "\<dots> \<le> sum (g \<circ> f) (insert x F)"
using insert by auto
finally show ?case .
qed
--- a/src/HOL/Groups_List.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Groups_List.thy Mon Oct 17 14:37:32 2016 +0200
@@ -73,15 +73,15 @@
from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp
qed
-sublocale setsum: comm_monoid_list_set plus 0
+sublocale sum: comm_monoid_list_set plus 0
rewrites
"monoid_list.F plus 0 = sum_list"
- and "comm_monoid_set.F plus 0 = setsum"
+ and "comm_monoid_set.F plus 0 = sum"
proof -
show "comm_monoid_list_set plus 0" ..
- then interpret setsum: comm_monoid_list_set plus 0 .
+ then interpret sum: comm_monoid_list_set plus 0 .
from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp
- from setsum_def show "comm_monoid_set.F plus 0 = setsum" by (auto intro: sym)
+ from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym)
qed
end
@@ -134,13 +134,13 @@
shows "sum_list (map f (filter P xs)) = sum_list (map f xs)"
using assms by (induct xs) auto
-lemma (in comm_monoid_add) distinct_sum_list_conv_Setsum:
- "distinct xs \<Longrightarrow> sum_list xs = Setsum (set xs)"
+lemma (in comm_monoid_add) distinct_sum_list_conv_Sum:
+ "distinct xs \<Longrightarrow> sum_list xs = Sum (set xs)"
by (induct xs) simp_all
lemma sum_list_upt[simp]:
"m \<le> n \<Longrightarrow> sum_list [m..<n] = \<Sum> {m..<n}"
-by(simp add: distinct_sum_list_conv_Setsum)
+by(simp add: distinct_sum_list_conv_Sum)
lemma sum_list_eq_0_nat_iff_nat [simp]:
"sum_list ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
@@ -200,25 +200,25 @@
shows "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sum>x\<leftarrow>xs. f x) \<le> (\<Sum>x\<leftarrow>xs. g x)"
by (induct xs) (simp, simp add: add_mono)
-lemma (in monoid_add) sum_list_distinct_conv_setsum_set:
- "distinct xs \<Longrightarrow> sum_list (map f xs) = setsum f (set xs)"
+lemma (in monoid_add) sum_list_distinct_conv_sum_set:
+ "distinct xs \<Longrightarrow> sum_list (map f xs) = sum f (set xs)"
by (induct xs) simp_all
-lemma (in monoid_add) interv_sum_list_conv_setsum_set_nat:
- "sum_list (map f [m..<n]) = setsum f (set [m..<n])"
- by (simp add: sum_list_distinct_conv_setsum_set)
+lemma (in monoid_add) interv_sum_list_conv_sum_set_nat:
+ "sum_list (map f [m..<n]) = sum f (set [m..<n])"
+ by (simp add: sum_list_distinct_conv_sum_set)
-lemma (in monoid_add) interv_sum_list_conv_setsum_set_int:
- "sum_list (map f [k..l]) = setsum f (set [k..l])"
- by (simp add: sum_list_distinct_conv_setsum_set)
+lemma (in monoid_add) interv_sum_list_conv_sum_set_int:
+ "sum_list (map f [k..l]) = sum f (set [k..l])"
+ by (simp add: sum_list_distinct_conv_sum_set)
-text \<open>General equivalence between @{const sum_list} and @{const setsum}\<close>
-lemma (in monoid_add) sum_list_setsum_nth:
+text \<open>General equivalence between @{const sum_list} and @{const sum}\<close>
+lemma (in monoid_add) sum_list_sum_nth:
"sum_list xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
- using interv_sum_list_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
+ using interv_sum_list_conv_sum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
-lemma sum_list_map_eq_setsum_count:
- "sum_list (map f xs) = setsum (\<lambda>x. count_list xs x * f x) (set xs)"
+lemma sum_list_map_eq_sum_count:
+ "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) (set xs)"
proof(induction xs)
case (Cons x xs)
show ?case (is "?l = ?r")
@@ -227,7 +227,7 @@
have "?l = f x + (\<Sum>x\<in>set xs. count_list xs x * f x)" by (simp add: Cons.IH)
also have "set xs = insert x (set xs - {x})" using \<open>x \<in> set xs\<close>by blast
also have "f x + (\<Sum>x\<in>insert x (set xs - {x}). count_list xs x * f x) = ?r"
- by (simp add: setsum.insert_remove eq_commute)
+ by (simp add: sum.insert_remove eq_commute)
finally show ?thesis .
next
assume "x \<notin> set xs"
@@ -236,17 +236,17 @@
qed
qed simp
-lemma sum_list_map_eq_setsum_count2:
+lemma sum_list_map_eq_sum_count2:
assumes "set xs \<subseteq> X" "finite X"
-shows "sum_list (map f xs) = setsum (\<lambda>x. count_list xs x * f x) X"
+shows "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) X"
proof-
let ?F = "\<lambda>x. count_list xs x * f x"
- have "setsum ?F X = setsum ?F (set xs \<union> (X - set xs))"
+ have "sum ?F X = sum ?F (set xs \<union> (X - set xs))"
using Un_absorb1[OF assms(1)] by(simp)
- also have "\<dots> = setsum ?F (set xs)"
+ also have "\<dots> = sum ?F (set xs)"
using assms(2)
- by(simp add: setsum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel)
- finally show ?thesis by(simp add:sum_list_map_eq_setsum_count)
+ by(simp add: sum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel)
+ finally show ?thesis by(simp add:sum_list_map_eq_sum_count)
qed
lemma sum_list_nonneg:
@@ -298,15 +298,15 @@
subsection \<open>Tools setup\<close>
-lemmas setsum_code = setsum.set_conv_list
+lemmas sum_code = sum.set_conv_list
-lemma setsum_set_upto_conv_sum_list_int [code_unfold]:
- "setsum f (set [i..j::int]) = sum_list (map f [i..j])"
- by (simp add: interv_sum_list_conv_setsum_set_int)
+lemma sum_set_upto_conv_sum_list_int [code_unfold]:
+ "sum f (set [i..j::int]) = sum_list (map f [i..j])"
+ by (simp add: interv_sum_list_conv_sum_set_int)
-lemma setsum_set_upt_conv_sum_list_nat [code_unfold]:
- "setsum f (set [m..<n]) = sum_list (map f [m..<n])"
- by (simp add: interv_sum_list_conv_setsum_set_nat)
+lemma sum_set_upt_conv_sum_list_nat [code_unfold]:
+ "sum f (set [m..<n]) = sum_list (map f [m..<n])"
+ by (simp add: interv_sum_list_conv_sum_set_nat)
lemma sum_list_transfer[transfer_rule]:
includes lifting_syntax
--- a/src/HOL/HOLCF/Universal.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/HOLCF/Universal.thy Mon Oct 17 14:37:32 2016 +0200
@@ -45,9 +45,9 @@
unfolding node_def less_Suc_eq_le set_encode_def
apply (rule order_trans [OF _ le_prod_encode_2])
apply (rule order_trans [OF _ le_prod_encode_2])
-apply (rule order_trans [where y="setsum (op ^ 2) {b}"])
+apply (rule order_trans [where y="sum (op ^ 2) {b}"])
apply (simp add: nat_less_power2 [THEN order_less_imp_le])
-apply (erule setsum_mono2, simp, simp)
+apply (erule sum_mono2, simp, simp)
done
lemma eq_prod_encode_pairI:
--- a/src/HOL/Hoare_Parallel/OG_Examples.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Hoare_Parallel/OG_Examples.thy Mon Oct 17 14:37:32 2016 +0200
@@ -506,10 +506,10 @@
lemma Example2_lemma2:
"!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
-apply(erule_tac t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
+apply(erule_tac t="sum (b(j := (Suc 0))) {0..<n}" in ssubst)
apply(frule_tac b=b in Example2_lemma2_aux)
-apply(erule_tac t="setsum b {0..<n}" in ssubst)
-apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
+apply(erule_tac t="sum b {0..<n}" in ssubst)
+apply(subgoal_tac "Suc (sum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(sum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
apply(rotate_tac -1)
apply(erule ssubst)
apply(subgoal_tac "j\<le>j")
@@ -534,9 +534,9 @@
COEND
\<lbrace>\<acute>x=n\<rbrace>"
apply oghoare
-apply (simp_all cong del: setsum.strong_cong)
+apply (simp_all cong del: sum.strong_cong)
apply (tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
-apply (simp_all cong del: setsum.strong_cong)
+apply (simp_all cong del: sum.strong_cong)
apply(erule (1) Example2_lemma2)
apply(erule (1) Example2_lemma2)
apply(erule (1) Example2_lemma2)
--- a/src/HOL/Hoare_Parallel/RG_Examples.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Mon Oct 17 14:37:32 2016 +0200
@@ -172,10 +172,10 @@
lemma Example2_lemma2:
"\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
-apply(erule_tac t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
+apply(erule_tac t="sum (b(j := (Suc 0))) {0..<n}" in ssubst)
apply(frule_tac b=b in Example2_lemma2_aux)
-apply(erule_tac t="setsum b {0..<n}" in ssubst)
-apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
+apply(erule_tac t="sum b {0..<n}" in ssubst)
+apply(subgoal_tac "Suc (sum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(sum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
apply(rotate_tac -1)
apply(erule ssubst)
apply(subgoal_tac "j\<le>j")
--- a/src/HOL/IMP/Abs_Int0.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/IMP/Abs_Int0.thy Mon Oct 17 14:37:32 2016 +0200
@@ -319,7 +319,7 @@
"m_s S X = (\<Sum> x \<in> X. m(S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"
-by(simp add: m_s_def) (metis mult.commute of_nat_id setsum_bounded_above[OF h])
+by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h])
fun m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
"m_o (Some S) X = m_s S X" |
@@ -336,9 +336,9 @@
proof-
let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)"
have "m_c C = (\<Sum>i<?a. m_o (annos C ! i) ?X)"
- by(simp add: m_c_def sum_list_setsum_nth atLeast0LessThan)
+ by(simp add: m_c_def sum_list_sum_nth atLeast0LessThan)
also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"
- apply(rule setsum_mono) using m_o_h[OF finite_Cvars] by simp
+ apply(rule sum_mono) using m_o_h[OF finite_Cvars] by simp
also have "\<dots> = ?a * (h * ?n + 1)" by simp
finally show ?thesis .
qed
@@ -409,7 +409,7 @@
from assms(2,3,4) have "EX x:X. S1 x < S2 x"
by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans)
hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2)
- from setsum_strict_mono_ex1[OF `finite X` 1 2]
+ from sum_strict_mono_ex1[OF `finite X` 1 2]
show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" .
qed
@@ -453,9 +453,9 @@
hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast
have "(\<Sum>i<size(annos C2). m_o (annos C2 ! i) ?X)
< (\<Sum>i<size(annos C2). m_o (annos C1 ! i) ?X)"
- apply(rule setsum_strict_mono_ex1) using 1 2 by (auto)
+ apply(rule sum_strict_mono_ex1) using 1 2 by (auto)
thus ?thesis
- by(simp add: m_c_def vars_acom_def strip_eq sum_list_setsum_nth atLeast0LessThan size_annos_same[OF strip_eq])
+ by(simp add: m_c_def vars_acom_def strip_eq sum_list_sum_nth atLeast0LessThan size_annos_same[OF strip_eq])
qed
end
--- a/src/HOL/IMP/Abs_Int1.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/IMP/Abs_Int1.thy Mon Oct 17 14:37:32 2016 +0200
@@ -109,7 +109,7 @@
"m_s S X = (\<Sum> x \<in> X. m(fun S x))"
lemma m_s_h: "finite X \<Longrightarrow> m_s S X \<le> h * card X"
-by(simp add: m_s_def) (metis mult.commute of_nat_id setsum_bounded_above[OF h])
+by(simp add: m_s_def) (metis mult.commute of_nat_id sum_bounded_above[OF h])
definition m_o :: "'av st option \<Rightarrow> vname set \<Rightarrow> nat" ("m\<^sub>o") where
"m_o opt X = (case opt of None \<Rightarrow> h * card X + 1 | Some S \<Rightarrow> m_s S X)"
@@ -125,9 +125,9 @@
proof-
let ?X = "vars C" let ?n = "card ?X" let ?a = "size(annos C)"
have "m_c C = (\<Sum>i<?a. m_o (annos C ! i) ?X)"
- by(simp add: m_c_def sum_list_setsum_nth atLeast0LessThan)
+ by(simp add: m_c_def sum_list_sum_nth atLeast0LessThan)
also have "\<dots> \<le> (\<Sum>i<?a. h * ?n + 1)"
- apply(rule setsum_mono) using m_o_h[OF finite_Cvars] by simp
+ apply(rule sum_mono) using m_o_h[OF finite_Cvars] by simp
also have "\<dots> = ?a * (h * ?n + 1)" by simp
finally show ?thesis .
qed
@@ -191,7 +191,7 @@
from assms(2,3,4) have "EX x:X. S1 x < S2 x"
by(simp add: fun_eq_iff) (metis Compl_iff le_neq_trans)
hence 2: "\<exists>x\<in>X. m(S1 x) > m(S2 x)" by (metis m2)
- from setsum_strict_mono_ex1[OF `finite X` 1 2]
+ from sum_strict_mono_ex1[OF `finite X` 1 2]
show "(\<Sum>x\<in>X. m (S2 x)) < (\<Sum>x\<in>X. m (S1 x))" .
qed
@@ -237,9 +237,9 @@
hence 2: "\<exists>i < size(annos C2). ?P i" using `i < size(annos C2)` by blast
have "(\<Sum>i<size(annos C2). m_o (annos C2 ! i) ?X)
< (\<Sum>i<size(annos C2). m_o (annos C1 ! i) ?X)"
- apply(rule setsum_strict_mono_ex1) using 1 2 by (auto)
+ apply(rule sum_strict_mono_ex1) using 1 2 by (auto)
thus ?thesis
- by(simp add: m_c_def vars_acom_def strip_eq sum_list_setsum_nth atLeast0LessThan size_annos_same[OF strip_eq])
+ by(simp add: m_c_def vars_acom_def strip_eq sum_list_sum_nth atLeast0LessThan size_annos_same[OF strip_eq])
qed
end
--- a/src/HOL/IMP/Abs_Int3.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/IMP/Abs_Int3.thy Mon Oct 17 14:37:32 2016 +0200
@@ -343,7 +343,7 @@
shows "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))"
proof-
from assms have "\<forall>x. m(S1 x) \<ge> m(S2 x)" by (metis m_anti_mono)
- thus "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))" by (metis setsum_mono)
+ thus "(\<Sum>x\<in>X. m (S2 x)) \<le> (\<Sum>x\<in>X. m (S1 x))" by (metis sum_mono)
qed
lemma m_s_anti_mono: "S1 \<le> S2 \<Longrightarrow> m_s S1 X \<ge> m_s S2 X"
@@ -361,7 +361,7 @@
by(auto simp add: Ball_def)
hence 2: "\<exists>x\<in>X. m(S1 x) > m(S1 x \<nabla> S2 x)"
using assms(3) m_widen by blast
- from setsum_strict_mono_ex1[OF `finite X` 1 2]
+ from sum_strict_mono_ex1[OF `finite X` 1 2]
show ?thesis .
qed
@@ -390,11 +390,11 @@
lemma m_c_widen:
"strip C1 = strip C2 \<Longrightarrow> top_on_acom C1 (-vars C1) \<Longrightarrow> top_on_acom C2 (-vars C2)
\<Longrightarrow> \<not> C2 \<le> C1 \<Longrightarrow> m_c (C1 \<nabla> C2) < m_c C1"
-apply(auto simp: m_c_def widen_acom_def map2_acom_def size_annos[symmetric] anno_def[symmetric]sum_list_setsum_nth)
+apply(auto simp: m_c_def widen_acom_def map2_acom_def size_annos[symmetric] anno_def[symmetric]sum_list_sum_nth)
apply(subgoal_tac "length(annos C2) = length(annos C1)")
prefer 2 apply (simp add: size_annos_same2)
apply (auto)
-apply(rule setsum_strict_mono_ex1)
+apply(rule sum_strict_mono_ex1)
apply(auto simp add: m_o_anti_mono vars_acom_def anno_def top_on_acom_def top_on_opt_widen widen1 less_eq_acom_def listrel_iff_nth)
apply(rule_tac x=p in bexI)
apply (auto simp: vars_acom_def m_o_widen top_on_acom_def)
@@ -415,7 +415,7 @@
hence 2: "\<exists>x\<in>X. n(S1 x \<triangle> S2 x) < n(S1 x)"
by (metis assms(3-5) eq_iff less_le_not_le n_narrow)
show ?thesis
- apply(rule setsum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+
+ apply(rule sum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+
qed
lemma n_s_narrow: "finite X \<Longrightarrow> fun S1 = fun S2 on -X \<Longrightarrow> S2 \<le> S1 \<Longrightarrow> S1 \<triangle> S2 < S1
@@ -446,12 +446,12 @@
lemma n_c_narrow: "strip C1 = strip C2
\<Longrightarrow> top_on_acom C1 (- vars C1) \<Longrightarrow> top_on_acom C2 (- vars C2)
\<Longrightarrow> C2 \<le> C1 \<Longrightarrow> C1 \<triangle> C2 < C1 \<Longrightarrow> n\<^sub>c (C1 \<triangle> C2) < n\<^sub>c C1"
-apply(auto simp: n_c_def narrow_acom_def sum_list_setsum_nth)
+apply(auto simp: n_c_def narrow_acom_def sum_list_sum_nth)
apply(subgoal_tac "length(annos C2) = length(annos C1)")
prefer 2 apply (simp add: size_annos_same2)
apply (auto)
apply(simp add: less_annos_iff le_iff_le_annos)
-apply(rule setsum_strict_mono_ex1)
+apply(rule sum_strict_mono_ex1)
apply (auto simp: vars_acom_def top_on_acom_def)
apply (metis n_o_narrow nth_mem finite_cvars less_imp_le le_less order_refl)
apply(rule_tac x=i in bexI)
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy Mon Oct 17 14:37:32 2016 +0200
@@ -203,9 +203,9 @@
by (fastforce simp: le_st_def lookup_def)
have "(\<Sum>y\<in>?Y'\<inter>?X. m(?g y)+1) < (\<Sum>y\<in>?Y'\<inter>?X. m(?f y)+1)"
using `u:?X` `u:?Y'` `m(?g u) < m(?f u)`
- by(fastforce intro!: setsum_strict_mono_ex1[OF _ 1])
+ by(fastforce intro!: sum_strict_mono_ex1[OF _ 1])
also have "\<dots> \<le> (\<Sum>y\<in>?X'. m(?f y)+1)"
- by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
+ by(simp add: sum_mono3[OF _ `?Y'\<inter>?X <= ?X'`])
finally show ?thesis .
next
assume "u \<notin> ?Y'"
@@ -217,11 +217,11 @@
qed
also have "\<dots> < (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) + (\<Sum>y\<in>{u}. m(?f y)+1)" by simp
also have "(\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?g y)+1) \<le> (\<Sum>y\<in>?Y'\<inter>?X-{u}. m(?f y)+1)"
- using 1 by(blast intro: setsum_mono)
+ using 1 by(blast intro: sum_mono)
also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
- by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
+ by(simp add: sum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
- using `u:?X'` by(subst setsum.union_disjoint[symmetric]) auto
+ using `u:?X'` by(subst sum.union_disjoint[symmetric]) auto
also have "\<dots> = (\<Sum>x\<in>?X'. m(?f x)+1)"
using `u : ?X'` by(simp add:insert_absorb)
finally show ?thesis by (blast intro: add_right_mono)
--- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy Mon Oct 17 14:37:32 2016 +0200
@@ -301,10 +301,10 @@
shows "m_st m_ivl S \<le> 2 * card X"
proof(auto simp: m_st_def)
have "(\<Sum>x\<in>set(dom S). m_ivl (fun S x)) \<le> (\<Sum>x\<in>set(dom S). 2)" (is "?L \<le> _")
- by(rule setsum_mono)(simp add:m_ivl_height)
+ by(rule sum_mono)(simp add:m_ivl_height)
also have "\<dots> \<le> (\<Sum>x\<in>X. 2)"
- by(rule setsum_mono3[OF assms]) simp
- also have "\<dots> = 2 * card X" by(simp add: setsum_constant)
+ by(rule sum_mono3[OF assms]) simp
+ also have "\<dots> = 2 * card X" by(simp add: sum_constant)
finally show "?L \<le> \<dots>" .
qed
@@ -319,13 +319,13 @@
have "(\<Sum>y\<in>?Y. m_ivl(?g y)) = (\<Sum>y\<in>(?Y-?X) \<union> (?Y\<inter>?X). m_ivl(?g y))"
by (metis Un_Diff_Int)
also have "\<dots> = (\<Sum>y\<in>?Y-?X. m_ivl(?g y)) + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))"
- by(subst setsum.union_disjoint) auto
+ by(subst sum.union_disjoint) auto
also have "(\<Sum>y\<in>?Y-?X. m_ivl(?g y)) = 0" using 0 by simp
also have "0 + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y)) = (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))" by simp
also have "\<dots> \<le> (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?f y))"
- by(rule setsum_mono)(simp add: 1)
+ by(rule sum_mono)(simp add: 1)
also have "\<dots> \<le> (\<Sum>y\<in>?X. m_ivl(?f y))"
- by(simp add: setsum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2])
+ by(simp add: sum_mono3[of "?X" "?Y Int ?X", OF _ Int_lower2])
finally show "(\<Sum>y\<in>?Y. m_ivl(?g y)) \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
by (metis add_less_cancel_left)
qed
@@ -340,7 +340,7 @@
proof cases
assume "x : ?Y"
have "?L < (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
- proof(rule setsum_strict_mono_ex1, simp)
+ proof(rule sum_strict_mono_ex1, simp)
show "\<forall>x\<in>?X\<inter>?Y. m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
by (metis m_ivl_anti_mono widen1)
next
@@ -348,12 +348,12 @@
using `x:?X` `x:?Y` `\<not> lookup S2 x \<sqsubseteq> ?f x`
by (metis IntI m_ivl_widen lookup_def)
qed
- also have "\<dots> \<le> ?R" by(simp add: setsum_mono3[OF _ Int_lower1])
+ also have "\<dots> \<le> ?R" by(simp add: sum_mono3[OF _ Int_lower1])
finally show ?thesis .
next
assume "x ~: ?Y"
have "?L \<le> (\<Sum>x\<in>?X\<inter>?Y. m_ivl(?f x))"
- proof(rule setsum_mono, simp)
+ proof(rule sum_mono, simp)
fix x assume "x:?X \<and> x:?Y" show "m_ivl(?f x \<nabla> ?g x) \<le> m_ivl (?f x)"
by (metis m_ivl_anti_mono widen1)
qed
@@ -363,7 +363,7 @@
also have "\<dots> = (\<Sum>y\<in>insert x (?X\<inter>?Y). m_ivl(?f y))"
using `x ~: ?Y` by simp
also have "\<dots> \<le> (\<Sum>x\<in>?X. m_ivl(?f x))"
- by(rule setsum_mono3)(insert `x:?X`, auto)
+ by(rule sum_mono3)(insert `x:?X`, auto)
finally show ?thesis .
qed
} with assms show ?thesis
@@ -376,7 +376,7 @@
shows "n_st n_ivl X S1 \<le> n_st n_ivl X S2"
proof-
have "(\<Sum>x\<in>X. n_ivl(lookup S1 x)) \<le> (\<Sum>x\<in>X. n_ivl(lookup S2 x))"
- apply(rule setsum_mono) using assms
+ apply(rule sum_mono) using assms
by(auto simp: le_st_def lookup_def n_ivl_mono split: if_splits)
thus ?thesis by(simp add: n_st_def)
qed
@@ -395,7 +395,7 @@
by(auto simp: le_st_def narrow_st_def lookup_def intro: n_ivl_narrow
split: if_splits)
have "(\<Sum>x\<in>X. n_ivl(lookup (S1 \<triangle> S2) x)) < (\<Sum>x\<in>X. n_ivl(lookup S1 x))"
- apply(rule setsum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+
+ apply(rule sum_strict_mono_ex1[OF `finite X`]) using 1 2 by blast+
thus ?thesis by(simp add: n_st_def)
qed
@@ -501,7 +501,7 @@
apply(subgoal_tac "length(annos c') = length(annos c)")
prefer 2 apply (simp add: size_annos_same2)
apply (auto)
-apply(rule setsum_strict_mono_ex1)
+apply(rule sum_strict_mono_ex1)
apply simp
apply (clarsimp)
apply(erule m_o_anti_mono)
@@ -522,7 +522,7 @@
apply(subgoal_tac "length(annos c') = length(annos c)")
prefer 2 apply (simp add: size_annos_same2)
apply (auto)
-apply(rule setsum_strict_mono_ex1)
+apply(rule sum_strict_mono_ex1)
apply simp
apply (clarsimp)
apply(rule n_o_mono)
--- a/src/HOL/Inequalities.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Inequalities.thy Mon Oct 17 14:37:32 2016 +0200
@@ -7,7 +7,7 @@
imports Real_Vector_Spaces
begin
-lemma Setsum_Icc_int: "(m::int) \<le> n \<Longrightarrow> \<Sum> {m..n} = (n*(n+1) - m*(m-1)) div 2"
+lemma Sum_Icc_int: "(m::int) \<le> n \<Longrightarrow> \<Sum> {m..n} = (n*(n+1) - m*(m-1)) div 2"
proof(induct i == "nat(n-m)" arbitrary: m n)
case 0
hence "m = n" by arith
@@ -25,26 +25,26 @@
finally show ?case .
qed
-lemma Setsum_Icc_nat: assumes "(m::nat) \<le> n"
+lemma Sum_Icc_nat: assumes "(m::nat) \<le> n"
shows "\<Sum> {m..n} = (n*(n+1) - m*(m-1)) div 2"
proof -
have "m*(m-1) \<le> n*(n + 1)"
using assms by (meson diff_le_self order_trans le_add1 mult_le_mono)
hence "int(\<Sum> {m..n}) = int((n*(n+1) - m*(m-1)) div 2)" using assms
- by (auto simp: Setsum_Icc_int[transferred, OF assms] zdiv_int of_nat_mult simp del: of_nat_setsum
+ by (auto simp: Sum_Icc_int[transferred, OF assms] zdiv_int of_nat_mult simp del: of_nat_sum
split: zdiff_int_split)
thus ?thesis
using of_nat_eq_iff by blast
qed
-lemma Setsum_Ico_nat: assumes "(m::nat) \<le> n"
+lemma Sum_Ico_nat: assumes "(m::nat) \<le> n"
shows "\<Sum> {m..<n} = (n*(n-1) - m*(m-1)) div 2"
proof cases
assume "m < n"
hence "{m..<n} = {m..n-1}" by auto
hence "\<Sum>{m..<n} = \<Sum>{m..n-1}" by simp
also have "\<dots> = (n*(n-1) - m*(m-1)) div 2"
- using assms \<open>m < n\<close> by (simp add: Setsum_Icc_nat mult.commute)
+ using assms \<open>m < n\<close> by (simp add: Sum_Icc_nat mult.commute)
finally show ?thesis .
next
assume "\<not> m < n" with assms show ?thesis by simp
@@ -59,13 +59,13 @@
let ?S = "(\<Sum>j=0..<n. (\<Sum>k=0..<n. (a j - a k) * (b j - b k)))"
have "2 * (of_nat n * (\<Sum>j=0..<n. (a j * b j)) - (\<Sum>j=0..<n. b j) * (\<Sum>k=0..<n. a k)) = ?S"
by (simp only: one_add_one[symmetric] algebra_simps)
- (simp add: algebra_simps setsum_subtractf setsum.distrib setsum.commute[of "\<lambda>i j. a i * b j"] setsum_distrib_left)
+ (simp add: algebra_simps sum_subtractf sum.distrib sum.commute[of "\<lambda>i j. a i * b j"] sum_distrib_left)
also
{ fix i j::nat assume "i<n" "j<n"
hence "a i - a j \<le> 0 \<and> b i - b j \<ge> 0 \<or> a i - a j \<ge> 0 \<and> b i - b j \<le> 0"
using assms by (cases "i \<le> j") (auto simp: algebra_simps)
} then have "?S \<le> 0"
- by (auto intro!: setsum_nonpos simp: mult_le_0_iff)
+ by (auto intro!: sum_nonpos simp: mult_le_0_iff)
finally show ?thesis by (simp add: algebra_simps)
qed
@@ -75,6 +75,6 @@
(\<And>i j. \<lbrakk> i\<le>j; j<n \<rbrakk> \<Longrightarrow> b i \<ge> b j) \<Longrightarrow>
n * (\<Sum>i=0..<n. a i * b i) \<le> (\<Sum>i=0..<n. a i) * (\<Sum>i=0..<n. b i)"
using Chebyshev_sum_upper[where 'a=real, of n a b]
-by (simp del: of_nat_mult of_nat_setsum add: of_nat_mult[symmetric] of_nat_setsum[symmetric])
+by (simp del: of_nat_mult of_nat_sum add: of_nat_mult[symmetric] of_nat_sum[symmetric])
end
--- a/src/HOL/Int.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Int.thy Mon Oct 17 14:37:32 2016 +0200
@@ -866,12 +866,12 @@
qed
-subsection \<open>@{term setsum} and @{term setprod}\<close>
+subsection \<open>@{term sum} and @{term setprod}\<close>
-lemma of_nat_setsum [simp]: "of_nat (setsum f A) = (\<Sum>x\<in>A. of_nat(f x))"
+lemma of_nat_sum [simp]: "of_nat (sum f A) = (\<Sum>x\<in>A. of_nat(f x))"
by (induct A rule: infinite_finite_induct) auto
-lemma of_int_setsum [simp]: "of_int (setsum f A) = (\<Sum>x\<in>A. of_int(f x))"
+lemma of_int_sum [simp]: "of_int (sum f A) = (\<Sum>x\<in>A. of_int(f x))"
by (induct A rule: infinite_finite_induct) auto
lemma of_nat_setprod [simp]: "of_nat (setprod f A) = (\<Prod>x\<in>A. of_nat(f x))"
@@ -880,7 +880,7 @@
lemma of_int_setprod [simp]: "of_int (setprod f A) = (\<Prod>x\<in>A. of_int(f x))"
by (induct A rule: infinite_finite_induct) auto
-lemmas int_setsum = of_nat_setsum [where 'a=int]
+lemmas int_sum = of_nat_sum [where 'a=int]
lemmas int_setprod = of_nat_setprod [where 'a=int]
--- a/src/HOL/Library/BigO.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/BigO.thy Mon Oct 17 14:37:32 2016 +0200
@@ -23,7 +23,7 @@
to be inessential.)
\<^item> We no longer use \<open>+\<close> as output syntax for \<open>+o\<close>
\<^item> Lemmas involving \<open>sumr\<close> have been replaced by more general lemmas
- involving `\<open>setsum\<close>.
+ involving `\<open>sum\<close>.
\<^item> The library has been expanded, with e.g.~support for expressions of
the form \<open>f < g + O(h)\<close>.
@@ -546,21 +546,21 @@
done
-subsection \<open>Setsum\<close>
+subsection \<open>Sum\<close>
-lemma bigo_setsum_main: "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y \<Longrightarrow>
+lemma bigo_sum_main: "\<forall>x. \<forall>y \<in> A x. 0 \<le> h x y \<Longrightarrow>
\<exists>c. \<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> \<le> c * h x y \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)"
apply (auto simp add: bigo_def)
apply (rule_tac x = "\<bar>c\<bar>" in exI)
apply (subst abs_of_nonneg) back back
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply force
- apply (subst setsum_distrib_left)
+ apply (subst sum_distrib_left)
apply (rule allI)
apply (rule order_trans)
- apply (rule setsum_abs)
- apply (rule setsum_mono)
+ apply (rule sum_abs)
+ apply (rule sum_mono)
apply (rule order_trans)
apply (drule spec)+
apply (drule bspec)+
@@ -572,24 +572,24 @@
apply force
done
-lemma bigo_setsum1: "\<forall>x y. 0 \<le> h x y \<Longrightarrow>
+lemma bigo_sum1: "\<forall>x y. 0 \<le> h x y \<Longrightarrow>
\<exists>c. \<forall>x y. \<bar>f x y\<bar> \<le> c * h x y \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)"
- apply (rule bigo_setsum_main)
+ apply (rule bigo_sum_main)
apply force
apply clarsimp
apply (rule_tac x = c in exI)
apply force
done
-lemma bigo_setsum2: "\<forall>y. 0 \<le> h y \<Longrightarrow>
+lemma bigo_sum2: "\<forall>y. 0 \<le> h y \<Longrightarrow>
\<exists>c. \<forall>y. \<bar>f y\<bar> \<le> c * (h y) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f y) =o O(\<lambda>x. \<Sum>y \<in> A x. h y)"
- by (rule bigo_setsum1) auto
+ by (rule bigo_sum1) auto
-lemma bigo_setsum3: "f =o O(h) \<Longrightarrow>
+lemma bigo_sum3: "f =o O(h) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)"
- apply (rule bigo_setsum1)
+ apply (rule bigo_sum1)
apply (rule allI)+
apply (rule abs_ge_zero)
apply (unfold bigo_def)
@@ -603,44 +603,44 @@
apply (rule abs_ge_zero)
done
-lemma bigo_setsum4: "f =o g +o O(h) \<Longrightarrow>
+lemma bigo_sum4: "f =o g +o O(h) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o
(\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)"
apply (rule set_minus_imp_plus)
apply (subst fun_diff_def)
- apply (subst setsum_subtractf [symmetric])
+ apply (subst sum_subtractf [symmetric])
apply (subst right_diff_distrib [symmetric])
- apply (rule bigo_setsum3)
+ apply (rule bigo_sum3)
apply (subst fun_diff_def [symmetric])
apply (erule set_plus_imp_minus)
done
-lemma bigo_setsum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow>
+lemma bigo_sum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow>
\<forall>x. 0 \<le> h x \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o
O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))"
apply (subgoal_tac "(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y)) =
(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h (k x y)\<bar>)")
apply (erule ssubst)
- apply (erule bigo_setsum3)
+ apply (erule bigo_sum3)
apply (rule ext)
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
apply (subst abs_of_nonneg)
apply auto
done
-lemma bigo_setsum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow>
+lemma bigo_sum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 \<le> l x y \<Longrightarrow>
\<forall>x. 0 \<le> h x \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. l x y * f (k x y)) =o
(\<lambda>x. \<Sum>y \<in> A x. l x y * g (k x y)) +o
O(\<lambda>x. \<Sum>y \<in> A x. l x y * h (k x y))"
apply (rule set_minus_imp_plus)
apply (subst fun_diff_def)
- apply (subst setsum_subtractf [symmetric])
+ apply (subst sum_subtractf [symmetric])
apply (subst right_diff_distrib [symmetric])
- apply (rule bigo_setsum5)
+ apply (rule bigo_sum5)
apply (subst fun_diff_def [symmetric])
apply (drule set_plus_imp_minus)
apply auto
--- a/src/HOL/Library/Bourbaki_Witt_Fixpoint.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Bourbaki_Witt_Fixpoint.thy Mon Oct 17 14:37:32 2016 +0200
@@ -402,4 +402,4 @@
shows "bourbaki_witt_fixpoint Sup {(x, y). x \<le> y} f"
by unfold_locales (auto simp: assms Sup_upper order_on_defs Field_def intro: refl_onI transI antisymI Sup_least)
-end
\ No newline at end of file
+end
--- a/src/HOL/Library/Combine_PER.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Combine_PER.thy Mon Oct 17 14:37:32 2016 +0200
@@ -49,4 +49,4 @@
elim: equivpE)
qed
-end
\ No newline at end of file
+end
--- a/src/HOL/Library/Extended_Nat.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Extended_Nat.thy Mon Oct 17 14:37:32 2016 +0200
@@ -17,7 +17,7 @@
begin
lemma sums_SUP[simp, intro]: "f sums (SUP n. \<Sum>i<n. f i)"
- unfolding sums_def by (intro LIMSEQ_SUP monoI setsum_mono2 zero_le) auto
+ unfolding sums_def by (intro LIMSEQ_SUP monoI sum_mono2 zero_le) auto
lemma suminf_eq_SUP: "suminf f = (SUP n. \<Sum>i<n. f i)"
using sums_SUP by (rule sums_unique[symmetric])
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Mon Oct 17 14:37:32 2016 +0200
@@ -40,11 +40,11 @@
moreover
{ fix k :: nat
have "(\<Sum>j<k + i. f j) = (\<Sum>j=i..<k + i. f j) + (\<Sum>j=0..<i. f j)"
- by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
+ by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong)
also have "(\<Sum>j=i..<k + i. f j) = (\<Sum>j\<in>(\<lambda>n. n + i)`{0..<k}. f j)"
unfolding image_add_atLeastLessThan by simp
finally have "(\<Sum>j<k + i. f j) = (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)"
- by (auto simp: inj_on_def atLeast0LessThan setsum.reindex) }
+ by (auto simp: inj_on_def atLeast0LessThan sum.reindex) }
ultimately have "(\<lambda>k. (\<Sum>n<k + i. f n)) \<longlonglongrightarrow> l + (\<Sum>j<i. f j)"
by simp
then show ?thesis
@@ -215,15 +215,15 @@
done
done
-lemma setsum_le_suminf:
+lemma sum_le_suminf:
fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add, linorder_topology}"
- shows "summable f \<Longrightarrow> finite I \<Longrightarrow> \<forall>m\<in>- I. 0 \<le> f m \<Longrightarrow> setsum f I \<le> suminf f"
+ shows "summable f \<Longrightarrow> finite I \<Longrightarrow> \<forall>m\<in>- I. 0 \<le> f m \<Longrightarrow> sum f I \<le> suminf f"
by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto
lemma suminf_eq_SUP_real:
assumes X: "summable X" "\<And>i. 0 \<le> X i" shows "suminf X = (SUP i. \<Sum>n<i. X n::real)"
by (intro LIMSEQ_unique[OF summable_LIMSEQ] X LIMSEQ_incseq_SUP)
- (auto intro!: bdd_aboveI2[where M="\<Sum>i. X i"] setsum_le_suminf X monoI setsum_mono3)
+ (auto intro!: bdd_aboveI2[where M="\<Sum>i. X i"] sum_le_suminf X monoI sum_mono3)
subsection \<open>Defining the extended non-negative reals\<close>
@@ -394,11 +394,11 @@
done
qed
-lemma setsum_enn2ereal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. enn2ereal (f i)) = enn2ereal (setsum f I)"
- by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq)
+lemma sum_enn2ereal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. enn2ereal (f i)) = enn2ereal (sum f I)"
+ by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq)
-lemma transfer_e2ennreal_setsum [transfer_rule]:
- "rel_fun (rel_fun op = pcr_ennreal) (rel_fun op = pcr_ennreal) setsum setsum"
+lemma transfer_e2ennreal_sum [transfer_rule]:
+ "rel_fun (rel_fun op = pcr_ennreal) (rel_fun op = pcr_ennreal) sum sum"
by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def)
lemma enn2ereal_of_nat[simp]: "enn2ereal (of_nat n) = ereal n"
@@ -522,12 +522,12 @@
shows "a + b = top \<longleftrightarrow> a = top \<or> b = top"
by transfer (auto simp: top_ereal_def)
-lemma ennreal_setsum_less_top[simp]:
+lemma ennreal_sum_less_top[simp]:
fixes f :: "'a \<Rightarrow> ennreal"
shows "finite I \<Longrightarrow> (\<Sum>i\<in>I. f i) < top \<longleftrightarrow> (\<forall>i\<in>I. f i < top)"
by (induction I rule: finite_induct) auto
-lemma ennreal_setsum_eq_top[simp]:
+lemma ennreal_sum_eq_top[simp]:
fixes f :: "'a \<Rightarrow> ennreal"
shows "finite I \<Longrightarrow> (\<Sum>i\<in>I. f i) = top \<longleftrightarrow> (\<exists>i\<in>I. f i = top)"
by (induction I rule: finite_induct) auto
@@ -953,8 +953,8 @@
"0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> ennreal (a + b) = ennreal a + ennreal b"
by (transfer fixing: a b) (auto simp: max_absorb2)
-lemma setsum_ennreal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. ennreal (f i)) = ennreal (setsum f I)"
- by (induction I rule: infinite_finite_induct) (auto simp: setsum_nonneg)
+lemma sum_ennreal[simp]: "(\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> (\<Sum>i\<in>I. ennreal (f i)) = ennreal (sum f I)"
+ by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg)
lemma sum_list_ennreal[simp]:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x \<ge> 0"
@@ -1368,10 +1368,10 @@
by transfer (auto intro!: sup_continuous_add)
lemma ennreal_suminf_lessD: "(\<Sum>i. f i :: ennreal) < x \<Longrightarrow> f i < x"
- using le_less_trans[OF setsum_le_suminf[OF summableI, of "{i}" f]] by simp
+ using le_less_trans[OF sum_le_suminf[OF summableI, of "{i}" f]] by simp
lemma sums_ennreal[simp]: "(\<And>i. 0 \<le> f i) \<Longrightarrow> 0 \<le> x \<Longrightarrow> (\<lambda>i. ennreal (f i)) sums ennreal x \<longleftrightarrow> f sums x"
- unfolding sums_def by (simp add: always_eventually setsum_nonneg)
+ unfolding sums_def by (simp add: always_eventually sum_nonneg)
lemma summable_suminf_not_top: "(\<And>i. 0 \<le> f i) \<Longrightarrow> (\<Sum>i. ennreal (f i)) \<noteq> top \<Longrightarrow> summable f"
using summable_sums[OF summableI, of "\<lambda>i. ennreal (f i)"]
@@ -1383,7 +1383,7 @@
by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums)
lemma sums_enn2ereal[simp]: "(\<lambda>i. enn2ereal (f i)) sums enn2ereal x \<longleftrightarrow> f sums x"
- unfolding sums_def by (simp add: always_eventually setsum_nonneg)
+ unfolding sums_def by (simp add: always_eventually sum_nonneg)
lemma suminf_enn2ereal[simp]: "(\<Sum>i. enn2ereal (f i)) = enn2ereal (suminf f)"
by (rule sums_unique[symmetric]) (simp add: summable_sums)
@@ -1529,12 +1529,12 @@
by transfer
(simp add: SUP_ereal_add incseq_def le_fun_def max_absorb2 SUP_upper2)
-lemma ennreal_SUP_setsum:
+lemma ennreal_SUP_sum:
fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
shows "(\<And>i. i \<in> I \<Longrightarrow> incseq (f i)) \<Longrightarrow> (SUP n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. SUP n. f i n)"
unfolding incseq_def
by transfer
- (simp add: SUP_ereal_setsum incseq_def SUP_upper2 max_absorb2 setsum_nonneg)
+ (simp add: SUP_ereal_sum incseq_def SUP_upper2 max_absorb2 sum_nonneg)
lemma ennreal_liminf_minus:
fixes f :: "nat \<Rightarrow> ennreal"
--- a/src/HOL/Library/Extended_Real.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Extended_Real.thy Mon Oct 17 14:37:32 2016 +0200
@@ -15,20 +15,20 @@
text \<open>This should be part of @{theory Extended_Nat} or @{theory Order_Continuity}, but then the
AFP-entry \<open>Jinja_Thread\<close> fails, as it does overload certain named from @{theory Complex_Main}.\<close>
-lemma incseq_setsumI2:
+lemma incseq_sumI2:
fixes f :: "'i \<Rightarrow> nat \<Rightarrow> 'a::ordered_comm_monoid_add"
shows "(\<And>n. n \<in> A \<Longrightarrow> mono (f n)) \<Longrightarrow> mono (\<lambda>i. \<Sum>n\<in>A. f n i)"
- unfolding incseq_def by (auto intro: setsum_mono)
-
-lemma incseq_setsumI:
+ unfolding incseq_def by (auto intro: sum_mono)
+
+lemma incseq_sumI:
fixes f :: "nat \<Rightarrow> 'a::ordered_comm_monoid_add"
assumes "\<And>i. 0 \<le> f i"
- shows "incseq (\<lambda>i. setsum f {..< i})"
+ shows "incseq (\<lambda>i. sum f {..< i})"
proof (intro incseq_SucI)
fix n
- have "setsum f {..< n} + 0 \<le> setsum f {..<n} + f n"
+ have "sum f {..< n} + 0 \<le> sum f {..<n} + f n"
using assms by (rule add_left_mono)
- then show "setsum f {..< n} \<le> setsum f {..< Suc n}"
+ then show "sum f {..< n} \<le> sum f {..< Suc n}"
by auto
qed
@@ -762,7 +762,7 @@
shows "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a + b"
using add_mono[of 0 a 0 b] by simp
-lemma setsum_ereal[simp]: "(\<Sum>x\<in>A. ereal (f x)) = ereal (\<Sum>x\<in>A. f x)"
+lemma sum_ereal[simp]: "(\<Sum>x\<in>A. ereal (f x)) = ereal (\<Sum>x\<in>A. f x)"
proof (cases "finite A")
case True
then show ?thesis by induct auto
@@ -774,11 +774,11 @@
lemma sum_list_ereal [simp]: "sum_list (map (\<lambda>x. ereal (f x)) xs) = ereal (sum_list (map f xs))"
by (induction xs) simp_all
-lemma setsum_Pinfty:
+lemma sum_Pinfty:
fixes f :: "'a \<Rightarrow> ereal"
shows "(\<Sum>x\<in>P. f x) = \<infinity> \<longleftrightarrow> finite P \<and> (\<exists>i\<in>P. f i = \<infinity>)"
proof safe
- assume *: "setsum f P = \<infinity>"
+ assume *: "sum f P = \<infinity>"
show "finite P"
proof (rule ccontr)
assume "\<not> finite P"
@@ -790,7 +790,7 @@
assume "\<not> ?thesis"
then have "\<And>i. i \<in> P \<Longrightarrow> f i \<noteq> \<infinity>"
by auto
- with \<open>finite P\<close> have "setsum f P \<noteq> \<infinity>"
+ with \<open>finite P\<close> have "sum f P \<noteq> \<infinity>"
by induct auto
with * show False
by auto
@@ -798,18 +798,18 @@
next
fix i
assume "finite P" and "i \<in> P" and "f i = \<infinity>"
- then show "setsum f P = \<infinity>"
+ then show "sum f P = \<infinity>"
proof induct
case (insert x A)
show ?case using insert by (cases "x = i") auto
qed simp
qed
-lemma setsum_Inf:
+lemma sum_Inf:
fixes f :: "'a \<Rightarrow> ereal"
- shows "\<bar>setsum f A\<bar> = \<infinity> \<longleftrightarrow> finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)"
+ shows "\<bar>sum f A\<bar> = \<infinity> \<longleftrightarrow> finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)"
proof
- assume *: "\<bar>setsum f A\<bar> = \<infinity>"
+ assume *: "\<bar>sum f A\<bar> = \<infinity>"
have "finite A"
by (rule ccontr) (insert *, auto)
moreover have "\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>"
@@ -827,18 +827,18 @@
assume "finite A \<and> (\<exists>i\<in>A. \<bar>f i\<bar> = \<infinity>)"
then obtain i where "finite A" "i \<in> A" and "\<bar>f i\<bar> = \<infinity>"
by auto
- then show "\<bar>setsum f A\<bar> = \<infinity>"
+ then show "\<bar>sum f A\<bar> = \<infinity>"
proof induct
case (insert j A)
then show ?case
- by (cases rule: ereal3_cases[of "f i" "f j" "setsum f A"]) auto
+ by (cases rule: ereal3_cases[of "f i" "f j" "sum f A"]) auto
qed simp
qed
-lemma setsum_real_of_ereal:
+lemma sum_real_of_ereal:
fixes f :: "'i \<Rightarrow> ereal"
assumes "\<And>x. x \<in> S \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
- shows "(\<Sum>x\<in>S. real_of_ereal (f x)) = real_of_ereal (setsum f S)"
+ shows "(\<Sum>x\<in>S. real_of_ereal (f x)) = real_of_ereal (sum f S)"
proof -
have "\<forall>x\<in>S. \<exists>r. f x = ereal r"
proof
@@ -852,17 +852,17 @@
by simp
qed
-lemma setsum_ereal_0:
+lemma sum_ereal_0:
fixes f :: "'a \<Rightarrow> ereal"
assumes "finite A"
and "\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i"
shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
proof
- assume "setsum f A = 0" with assms show "\<forall>i\<in>A. f i = 0"
+ assume "sum f A = 0" with assms show "\<forall>i\<in>A. f i = 0"
proof (induction A)
case (insert a A)
then have "f a = 0 \<and> (\<Sum>a\<in>A. f a) = 0"
- by (subst ereal_add_nonneg_eq_0_iff[symmetric]) (simp_all add: setsum_nonneg)
+ by (subst ereal_add_nonneg_eq_0_iff[symmetric]) (simp_all add: sum_nonneg)
with insert show ?case
by simp
qed simp
@@ -1089,18 +1089,18 @@
"c \<ge> 0 \<Longrightarrow> (x + y) * ereal c = x * ereal c + y * ereal c"
by(cases x y rule: ereal2_cases)(simp_all add: ring_distribs)
-lemma setsum_ereal_right_distrib:
+lemma sum_ereal_right_distrib:
fixes f :: "'a \<Rightarrow> ereal"
- shows "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> r * setsum f A = (\<Sum>n\<in>A. r * f n)"
- by (induct A rule: infinite_finite_induct) (auto simp: ereal_right_distrib setsum_nonneg)
-
-lemma setsum_ereal_left_distrib:
- "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> setsum f A * r = (\<Sum>n\<in>A. f n * r :: ereal)"
- using setsum_ereal_right_distrib[of A f r] by (simp add: mult_ac)
-
-lemma setsum_distrib_right_ereal:
- "c \<ge> 0 \<Longrightarrow> setsum f A * ereal c = (\<Sum>x\<in>A. f x * c :: ereal)"
-by(subst setsum_comp_morphism[where h="\<lambda>x. x * ereal c", symmetric])(simp_all add: distrib_left_ereal_nn)
+ shows "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> r * sum f A = (\<Sum>n\<in>A. r * f n)"
+ by (induct A rule: infinite_finite_induct) (auto simp: ereal_right_distrib sum_nonneg)
+
+lemma sum_ereal_left_distrib:
+ "(\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> sum f A * r = (\<Sum>n\<in>A. f n * r :: ereal)"
+ using sum_ereal_right_distrib[of A f r] by (simp add: mult_ac)
+
+lemma sum_distrib_right_ereal:
+ "c \<ge> 0 \<Longrightarrow> sum f A * ereal c = (\<Sum>x\<in>A. f x * c :: ereal)"
+by(subst sum_comp_morphism[where h="\<lambda>x. x * ereal c", symmetric])(simp_all add: distrib_left_ereal_nn)
lemma ereal_le_epsilon:
fixes x y :: ereal
@@ -2182,7 +2182,7 @@
using pos[of i] by auto
qed
-lemma SUP_ereal_setsum:
+lemma SUP_ereal_sum:
fixes f g :: "'a \<Rightarrow> nat \<Rightarrow> ereal"
assumes "\<And>n. n \<in> A \<Longrightarrow> incseq (f n)"
and pos: "\<And>n i. n \<in> A \<Longrightarrow> 0 \<le> f n i"
@@ -2190,7 +2190,7 @@
proof (cases "finite A")
case True
then show ?thesis using assms
- by induct (auto simp: incseq_setsumI2 setsum_nonneg SUP_ereal_add_pos)
+ by induct (auto simp: incseq_sumI2 sum_nonneg SUP_ereal_add_pos)
next
case False
then show ?thesis by simp
@@ -3111,7 +3111,7 @@
have "summable f" using pos[THEN summable_ereal_pos] .
then show "(\<lambda>N. \<Sum>n<N. f n) \<longlonglongrightarrow> suminf f"
by (auto dest!: summable_sums simp: sums_def atLeast0LessThan)
- show "\<forall>n\<ge>0. setsum f {..<n} \<le> x"
+ show "\<forall>n\<ge>0. sum f {..<n} \<le> x"
using assms by auto
qed
@@ -3157,12 +3157,12 @@
have "0 \<le> g N"
using assms(2,1)[of N] by auto
}
- have "setsum f {..<n} \<le> setsum g {..<n}"
- using assms by (auto intro: setsum_mono)
+ have "sum f {..<n} \<le> sum g {..<n}"
+ using assms by (auto intro: sum_mono)
also have "\<dots> \<le> suminf g"
using \<open>\<And>N. 0 \<le> g N\<close>
by (rule suminf_upper)
- finally show "setsum f {..<n} \<le> suminf g" .
+ finally show "sum f {..<n} \<le> suminf g" .
qed (rule assms(2))
lemma suminf_half_series_ereal: "(\<Sum>n. (1/2 :: ereal) ^ Suc n) = 1"
@@ -3175,8 +3175,8 @@
and "\<And>i. 0 \<le> g i"
shows "(\<Sum>i. f i + g i) = suminf f + suminf g"
apply (subst (1 2 3) suminf_ereal_eq_SUP)
- unfolding setsum.distrib
- apply (intro assms ereal_add_nonneg_nonneg SUP_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+
+ unfolding sum.distrib
+ apply (intro assms ereal_add_nonneg_nonneg SUP_ereal_add_pos incseq_sumI sum_nonneg ballI)+
done
lemma suminf_cmult_ereal:
@@ -3184,8 +3184,8 @@
assumes "\<And>i. 0 \<le> f i"
and "0 \<le> a"
shows "(\<Sum>i. a * f i) = a * suminf f"
- by (auto simp: setsum_ereal_right_distrib[symmetric] assms
- ereal_zero_le_0_iff setsum_nonneg suminf_ereal_eq_SUP
+ by (auto simp: sum_ereal_right_distrib[symmetric] assms
+ ereal_zero_le_0_iff sum_nonneg suminf_ereal_eq_SUP
intro!: SUP_ereal_mult_left)
lemma suminf_PInfty:
@@ -3198,7 +3198,7 @@
have "(\<Sum>i<Suc i. f i) \<noteq> \<infinity>"
by auto
then show ?thesis
- unfolding setsum_Pinfty by simp
+ unfolding sum_Pinfty by simp
qed
lemma suminf_PInfty_fun:
@@ -3325,7 +3325,7 @@
fix n :: nat
have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
using assms
- by (auto intro!: SUP_ereal_setsum [symmetric])
+ by (auto intro!: SUP_ereal_sum [symmetric])
}
note * = this
show ?thesis
@@ -3338,7 +3338,7 @@
done
qed
-lemma suminf_setsum_ereal:
+lemma suminf_sum_ereal:
fixes f :: "_ \<Rightarrow> _ \<Rightarrow> ereal"
assumes nonneg: "\<And>i a. a \<in> A \<Longrightarrow> 0 \<le> f i a"
shows "(\<Sum>i. \<Sum>a\<in>A. f i a) = (\<Sum>a\<in>A. \<Sum>i. f i a)"
@@ -3346,7 +3346,7 @@
case True
then show ?thesis
using nonneg
- by induct (simp_all add: suminf_add_ereal setsum_nonneg)
+ by induct (simp_all add: suminf_add_ereal sum_nonneg)
next
case False
then show ?thesis by simp
@@ -3392,10 +3392,10 @@
have "(\<Sum>i<n. f (i + k)) = (\<Sum>i<n. (f \<circ> (\<lambda>i. i + k)) i)"
by simp
also have "\<dots> = (\<Sum>i\<in>(\<lambda>i. i + k) ` {..<n}. f i)"
- by (subst setsum.reindex) auto
- also have "\<dots> \<le> setsum f {..<n + k}"
- by (intro setsum_mono3) (auto simp: f)
- finally show "(\<Sum>i<n. f (i + k)) \<le> setsum f {..<n + k}" .
+ by (subst sum.reindex) auto
+ also have "\<dots> \<le> sum f {..<n + k}"
+ by (intro sum_mono3) (auto simp: f)
+ finally show "(\<Sum>i<n. f (i + k)) \<le> sum f {..<n + k}" .
qed
qed
@@ -3445,7 +3445,7 @@
qed
qed
-lemma SUP_ereal_setsum_directed:
+lemma SUP_ereal_sum_directed:
fixes f g :: "'a \<Rightarrow> 'b \<Rightarrow> ereal"
assumes "I \<noteq> {}"
assumes directed: "\<And>N i j. N \<subseteq> A \<Longrightarrow> i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> \<exists>k\<in>I. \<forall>n\<in>N. f n i \<le> f n k \<and> f n j \<le> f n k"
@@ -3458,12 +3458,12 @@
moreover have "(SUP i:I. f n i + (\<Sum>l\<in>N. f l i)) = (SUP i:I. f n i) + (SUP i:I. \<Sum>l\<in>N. f l i)"
proof (rule SUP_ereal_add_directed)
fix i assume "i \<in> I" then show "0 \<le> f n i" "0 \<le> (\<Sum>l\<in>N. f l i)"
- using insert by (auto intro!: setsum_nonneg nonneg)
+ using insert by (auto intro!: sum_nonneg nonneg)
next
fix i j assume "i \<in> I" "j \<in> I"
from directed[OF \<open>insert n N \<subseteq> A\<close> this] guess k ..
then show "\<exists>k\<in>I. f n i + (\<Sum>l\<in>N. f l j) \<le> f n k + (\<Sum>l\<in>N. f l k)"
- by (intro bexI[of _ k]) (auto intro!: ereal_add_mono setsum_mono)
+ by (intro bexI[of _ k]) (auto intro!: ereal_add_mono sum_mono)
qed
ultimately show ?case
by simp
@@ -3482,7 +3482,7 @@
using \<open>I \<noteq> {}\<close> nonneg by (auto intro: SUP_upper2)
show "(SUP n. \<Sum>i<n. SUP n:I. f n i) = (SUP n:I. SUP j. \<Sum>i<j. f n i)"
apply (subst SUP_commute)
- apply (subst SUP_ereal_setsum_directed)
+ apply (subst SUP_ereal_sum_directed)
apply (auto intro!: assms simp: finite_subset)
done
qed
--- a/src/HOL/Library/FSet.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/FSet.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1011,7 +1011,7 @@
subsection \<open>Size setup\<close>
context includes fset.lifting begin
-lift_definition size_fset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a fset \<Rightarrow> nat" is "\<lambda>f. setsum (Suc \<circ> f)" .
+lift_definition size_fset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a fset \<Rightarrow> nat" is "\<lambda>f. sum (Suc \<circ> f)" .
end
instantiation fset :: (type) size begin
@@ -1030,7 +1030,7 @@
lemma fset_size_o_map: "inj f \<Longrightarrow> size_fset g \<circ> fimage f = size_fset (g \<circ> f)"
apply (subst fun_eq_iff)
- including fset.lifting by transfer (auto intro: setsum.reindex_cong subset_inj_on)
+ including fset.lifting by transfer (auto intro: sum.reindex_cong subset_inj_on)
setup \<open>
BNF_LFP_Size.register_size_global @{type_name fset} @{const_name size_fset}
--- a/src/HOL/Library/Finite_Map.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Finite_Map.thy Mon Oct 17 14:37:32 2016 +0200
@@ -831,4 +831,4 @@
lifting_update fmap.lifting
lifting_forget fmap.lifting
-end
\ No newline at end of file
+end
--- a/src/HOL/Library/Formal_Power_Series.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Formal_Power_Series.thy Mon Oct 17 14:37:32 2016 +0200
@@ -129,7 +129,7 @@
and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
shows "(\<Sum>j=0..k. \<Sum>i=0..j. f i (j - i) (n - j)) =
(\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
- by (induct k) (simp_all add: Suc_diff_le setsum.distrib add.assoc)
+ by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc)
instance fps :: (semiring_0) semigroup_mult
proof
@@ -141,7 +141,7 @@
(\<Sum>j=0..n. \<Sum>i=0..n - j. a$j * b$i * c$(n - j - i))"
by (rule fps_mult_assoc_lemma)
then show "((a * b) * c) $ n = (a * (b * c)) $ n"
- by (simp add: fps_mult_nth setsum_distrib_left setsum_distrib_right mult.assoc)
+ by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc)
qed
qed
@@ -149,7 +149,7 @@
fixes n :: nat
and f :: "nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
- by (rule setsum.reindex_bij_witness[where i="op - n" and j="op - n"]) auto
+ by (rule sum.reindex_bij_witness[where i="op - n" and j="op - n"]) auto
instance fps :: (comm_semiring_0) ab_semigroup_mult
proof
@@ -181,9 +181,9 @@
proof
fix a :: "'a fps"
show "1 * a = a"
- by (simp add: fps_ext fps_mult_nth mult_delta_left setsum.delta)
+ by (simp add: fps_ext fps_mult_nth mult_delta_left sum.delta)
show "a * 1 = a"
- by (simp add: fps_ext fps_mult_nth mult_delta_right setsum.delta')
+ by (simp add: fps_ext fps_mult_nth mult_delta_right sum.delta')
qed
instance fps :: (cancel_semigroup_add) cancel_semigroup_add
@@ -227,9 +227,9 @@
proof
fix a b c :: "'a fps"
show "(a + b) * c = a * c + b * c"
- by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum.distrib)
+ by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib)
show "a * (b + c) = a * b + a * c"
- by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum.distrib)
+ by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib)
qed
instance fps :: (semiring_0) semiring_0
@@ -276,7 +276,7 @@
lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
by (rule expand_fps_eq)
-lemma fps_setsum_nth: "setsum f S $ n = setsum (\<lambda>k. (f k) $ n) S"
+lemma fps_sum_nth: "sum f S $ n = sum (\<lambda>k. (f k) $ n) S"
proof (cases "finite S")
case True
then show ?thesis by (induct set: finite) auto
@@ -309,7 +309,7 @@
by (simp add: fps_ext)
lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
- by (simp add: fps_eq_iff fps_mult_nth setsum.neutral)
+ by (simp add: fps_eq_iff fps_mult_nth sum.neutral)
lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
@@ -321,17 +321,17 @@
lemma fps_const_mult_left: "fps_const (c::'a::semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
unfolding fps_eq_iff fps_mult_nth
- by (simp add: fps_const_def mult_delta_left setsum.delta)
+ by (simp add: fps_const_def mult_delta_left sum.delta)
lemma fps_const_mult_right: "f * fps_const (c::'a::semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
unfolding fps_eq_iff fps_mult_nth
- by (simp add: fps_const_def mult_delta_right setsum.delta')
+ by (simp add: fps_const_def mult_delta_right sum.delta')
lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
- by (simp add: fps_mult_nth mult_delta_left setsum.delta)
+ by (simp add: fps_mult_nth mult_delta_left sum.delta)
lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
- by (simp add: fps_mult_nth mult_delta_right setsum.delta')
+ by (simp add: fps_mult_nth mult_delta_right sum.delta')
subsection \<open>Formal power series form an integral domain\<close>
@@ -354,9 +354,9 @@
have "(a * b) $ (i + j) = (\<Sum>k=0..i+j. a $ k * b $ (i + j - k))"
by (rule fps_mult_nth)
also have "\<dots> = (a $ i * b $ (i + j - i)) + (\<Sum>k\<in>{0..i+j} - {i}. a $ k * b $ (i + j - k))"
- by (rule setsum.remove) simp_all
+ by (rule sum.remove) simp_all
also have "(\<Sum>k\<in>{0..i+j}-{i}. a $ k * b $ (i + j - k)) = 0"
- proof (rule setsum.neutral [rule_format])
+ proof (rule sum.neutral [rule_format])
fix k assume "k \<in> {0..i+j} - {i}"
then have "k < i \<or> i+j-k < j"
by auto
@@ -409,7 +409,7 @@
have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
by (simp add: fps_mult_nth)
also have "\<dots> = f $ (n - 1)"
- using False by (simp add: X_def mult_delta_left setsum.delta)
+ using False by (simp add: X_def mult_delta_left sum.delta)
finally show ?thesis
using False by simp
next
@@ -424,8 +424,8 @@
have "(a * X) $ n = (\<Sum>i = 0..n. a $ i * (if n - i = Suc 0 then 1 else 0))"
by (simp add: fps_times_def X_def)
also have "\<dots> = (\<Sum>i = 0..n. if i = n - 1 then if n = 0 then 0 else a $ i else 0)"
- by (intro setsum.cong) auto
- also have "\<dots> = (if n = 0 then 0 else a $ (n - 1))" by (simp add: setsum.delta)
+ by (intro sum.cong) auto
+ also have "\<dots> = (if n = 0 then 0 else a $ (n - 1))" by (simp add: sum.delta)
finally show ?thesis .
qed
@@ -587,13 +587,13 @@
have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))"
by (simp add: fps_mult_nth)
also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
- proof (intro setsum.cong)
+ proof (intro sum.cong)
fix x assume x: "x \<in> {0..?n}"
hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
by (elim disjE conjE) auto
qed auto
- also have "... = f $ subdegree f * g $ subdegree g" by (simp add: setsum.delta)
+ also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
finally show ?thesis .
qed
@@ -604,20 +604,20 @@
let ?n = "subdegree f + subdegree g"
have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth)
also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
- proof (intro setsum.cong)
+ proof (intro sum.cong)
fix x assume x: "x \<in> {0..?n}"
hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
by (elim disjE conjE) auto
qed auto
- also have "... = f $ subdegree f * g $ subdegree g" by (simp add: setsum.delta)
+ also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
also from assms have "... \<noteq> 0" by auto
finally show "(f * g) $ (subdegree f + subdegree g) \<noteq> 0" .
next
fix m assume m: "m < subdegree f + subdegree g"
have "(f * g) $ m = (\<Sum>i=0..m. f$i * g$(m-i))" by (simp add: fps_mult_nth)
also have "... = (\<Sum>i=0..m. 0)"
- proof (rule setsum.cong)
+ proof (rule sum.cong)
fix i assume "i \<in> {0..m}"
with m have "i < subdegree f \<or> m - i < subdegree g" by auto
thus "f$i * g$(m-i) = 0" by (elim disjE) auto
@@ -914,13 +914,13 @@
using kp by blast
qed
-lemma fps_sum_rep_nth: "(setsum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
+lemma fps_sum_rep_nth: "(sum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
(if n \<le> m then a$n else 0::'a::comm_ring_1)"
- apply (auto simp add: fps_setsum_nth cond_value_iff cong del: if_weak_cong)
- apply (simp add: setsum.delta')
+ apply (auto simp add: fps_sum_nth cond_value_iff cong del: if_weak_cong)
+ apply (simp add: sum.delta')
done
-lemma fps_notation: "(\<lambda>n. setsum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) \<longlonglongrightarrow> a"
+lemma fps_notation: "(\<lambda>n. sum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) \<longlonglongrightarrow> a"
(is "?s \<longlonglongrightarrow> a")
proof -
have "\<exists>n0. \<forall>n \<ge> n0. dist (?s n) a < r" if "r > 0" for r
@@ -964,7 +964,7 @@
subsection \<open>Inverses of formal power series\<close>
-declare setsum.cong[fundef_cong]
+declare sum.cong[fundef_cong]
instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse
begin
@@ -972,7 +972,7 @@
fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
where
"natfun_inverse f 0 = inverse (f$0)"
-| "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
+| "natfun_inverse f n = - inverse (f$0) * sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
definition fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))"
@@ -1012,10 +1012,10 @@
have d: "{0} \<inter> {1 .. n} = {}"
by auto
from f0 np have th0: "- (inverse f $ n) =
- (setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
+ (sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
by (cases n) (simp_all add: divide_inverse fps_inverse_def)
from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
- have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n"
+ have th1: "sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n"
by (simp add: field_simps)
have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
unfolding fps_mult_nth ifn ..
@@ -1077,7 +1077,7 @@
lemma fps_inverse_eq_0: "f$0 = 0 \<Longrightarrow> inverse (f :: 'a :: division_ring fps) = 0"
by simp
-lemma setsum_zero_lemma:
+lemma sum_zero_lemma:
fixes n::nat
assumes "0 < n"
shows "(\<Sum>i = 0..n. if n = i then 1 else if n - i = 1 then - 1 else 0) = (0::'a::field)"
@@ -1085,10 +1085,10 @@
let ?f = "\<lambda>i. if n = i then 1 else if n - i = 1 then - 1 else 0"
let ?g = "\<lambda>i. if i = n then 1 else if i = n - 1 then - 1 else 0"
let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
- have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
- by (rule setsum.cong) auto
- have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
- apply (rule setsum.cong)
+ have th1: "sum ?f {0..n} = sum ?g {0..n}"
+ by (rule sum.cong) auto
+ have th2: "sum ?g {0..n - 1} = sum ?h {0..n - 1}"
+ apply (rule sum.cong)
using assms
apply auto
done
@@ -1100,9 +1100,9 @@
by auto
show ?thesis
unfolding th1
- apply (simp add: setsum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
+ apply (simp add: sum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
unfolding th2
- apply (simp add: setsum.delta)
+ apply (simp add: sum.delta)
done
qed
@@ -1127,7 +1127,7 @@
lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) =
Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
apply (rule fps_inverse_unique)
- apply (simp_all add: fps_eq_iff fps_mult_nth setsum_zero_lemma)
+ apply (simp_all add: fps_eq_iff fps_mult_nth sum_zero_lemma)
done
lemma subdegree_inverse [simp]: "subdegree (inverse (f::'a::field fps)) = 0"
@@ -1546,24 +1546,24 @@
of_nat (i+1)* f $ (i+1) * g $ (n - i)"
let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
of_nat i* f $ i * g $ ((n + 1) - i)"
- have s0: "setsum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
- setsum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
- by (rule setsum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
- have s1: "setsum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
- setsum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
- by (rule setsum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
+ have s0: "sum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
+ sum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
+ by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
+ have s1: "sum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
+ sum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
+ by (rule sum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
by (simp only: mult.commute)
also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
- by (simp add: fps_mult_nth setsum.distrib[symmetric])
- also have "\<dots> = setsum ?h {0..n+1}"
- by (rule setsum.reindex_bij_witness_not_neutral
+ by (simp add: fps_mult_nth sum.distrib[symmetric])
+ also have "\<dots> = sum ?h {0..n+1}"
+ by (rule sum.reindex_bij_witness_not_neutral
[where S'="{}" and T'="{0}" and j="Suc" and i="\<lambda>i. i - 1"]) auto
also have "\<dots> = (fps_deriv (f * g)) $ n"
- apply (simp only: fps_deriv_nth fps_mult_nth setsum.distrib)
+ apply (simp only: fps_deriv_nth fps_mult_nth sum.distrib)
unfolding s0 s1
- unfolding setsum.distrib[symmetric] setsum_distrib_left
- apply (rule setsum.cong)
+ unfolding sum.distrib[symmetric] sum_distrib_left
+ apply (rule sum.cong)
apply (auto simp add: of_nat_diff field_simps)
done
finally show ?thesis .
@@ -1604,8 +1604,8 @@
"fps_deriv (f * fps_const (c::'a::comm_ring_1)) = fps_deriv f * fps_const c"
by simp
-lemma fps_deriv_setsum:
- "fps_deriv (setsum f S) = setsum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
+lemma fps_deriv_sum:
+ "fps_deriv (sum f S) = sum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
proof (cases "finite S")
case False
then show ?thesis by simp
@@ -1699,8 +1699,8 @@
"fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult.commute)
-lemma fps_nth_deriv_setsum:
- "fps_nth_deriv n (setsum f S) = setsum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
+lemma fps_nth_deriv_sum:
+ "fps_nth_deriv n (sum f S) = sum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
proof (cases "finite S")
case True
show ?thesis by (induct rule: finite_induct [OF True]) simp_all
@@ -1771,7 +1771,7 @@
also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
by (simp add: fps_mult_nth)
also have "\<dots> = 0"
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply auto
apply (case_tac "x = m")
using a0 apply simp
@@ -1784,11 +1784,11 @@
qed
qed
-lemma startsby_zero_setsum_depends:
+lemma startsby_zero_sum_depends:
assumes a0: "a $0 = (0::'a::idom)"
and kn: "n \<ge> k"
- shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
- apply (rule setsum.mono_neutral_right)
+ shows "sum (\<lambda>i. (a ^ i)$k) {0 .. n} = sum (\<lambda>i. (a ^ i)$k) {0 .. k}"
+ apply (rule sum.mono_neutral_right)
using kn
apply auto
apply (rule startsby_zero_power_prefix[rule_format, OF a0])
@@ -1805,10 +1805,10 @@
case (Suc n)
have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)"
by (simp add: field_simps)
- also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
+ also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
by (simp add: fps_mult_nth)
- also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
- apply (rule setsum.mono_neutral_right)
+ also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
+ apply (rule sum.mono_neutral_right)
apply simp
apply clarsimp
apply clarsimp
@@ -1946,19 +1946,19 @@
subsection \<open>Composition of FPSs\<close>
definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55)
- where "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
-
-lemma fps_compose_nth: "(a oo b)$n = setsum (\<lambda>i. a$i * (b^i$n)) {0..n}"
+ where "a oo b = Abs_fps (\<lambda>n. sum (\<lambda>i. a$i * (b^i$n)) {0..n})"
+
+lemma fps_compose_nth: "(a oo b)$n = sum (\<lambda>i. a$i * (b^i$n)) {0..n}"
by (simp add: fps_compose_def)
lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0"
by (simp add: fps_compose_nth)
lemma fps_compose_X[simp]: "a oo X = (a :: 'a::comm_ring_1 fps)"
- by (simp add: fps_ext fps_compose_def mult_delta_right setsum.delta')
+ by (simp add: fps_ext fps_compose_def mult_delta_right sum.delta')
lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
- by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
+ by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k"
unfolding numeral_fps_const by simp
@@ -1967,17 +1967,17 @@
unfolding neg_numeral_fps_const by simp
lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: 'a::comm_ring_1 fps)"
- by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum.delta not_le)
+ by (simp add: fps_eq_iff fps_compose_def mult_delta_left sum.delta not_le)
subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
subsubsection \<open>Rule 1\<close>
- (* {a_{n+k}}_0^infty Corresponds to (f - setsum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
+ (* {a_{n+k}}_0^infty Corresponds to (f - sum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
lemma fps_power_mult_eq_shift:
"X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
- Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * X^i) {0 .. k}"
+ Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * X^i) {0 .. k}"
(is "?lhs = ?rhs")
proof -
have "?lhs $ n = ?rhs $ n" for n :: nat
@@ -1988,11 +1988,11 @@
proof (induct k)
case 0
then show ?case
- by (simp add: fps_setsum_nth)
+ by (simp add: fps_sum_nth)
next
case (Suc k)
- have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
- (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} -
+ have "(Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
+ (Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} -
fps_const (a (Suc k)) * X^ Suc k) $ n"
by (simp add: field_simps)
also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
@@ -2051,10 +2051,10 @@
subsubsection \<open>Rule 5 --- summation and "division" by (1 - X)\<close>
-lemma fps_divide_X_minus1_setsum_lemma:
- "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
+lemma fps_divide_X_minus1_sum_lemma:
+ "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
proof -
- let ?sa = "Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
+ let ?sa = "Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)"
by simp
have "a$n = ((1 - X) * ?sa) $ n" for n
@@ -2071,14 +2071,14 @@
using False by simp_all
have f: "finite {0}" "finite {1}" "finite {2 .. n}"
"finite {0 .. n - 1}" "finite {n}" by simp_all
- have "((1 - X) * ?sa) $ n = setsum (\<lambda>i. (1 - X)$ i * ?sa $ (n - i)) {0 .. n}"
+ have "((1 - X) * ?sa) $ n = sum (\<lambda>i. (1 - X)$ i * ?sa $ (n - i)) {0 .. n}"
by (simp add: fps_mult_nth)
also have "\<dots> = a$n"
unfolding th0
- unfolding setsum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
- unfolding setsum.union_disjoint[OF f(2) f(3) d(2)]
+ unfolding sum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
+ unfolding sum.union_disjoint[OF f(2) f(3) d(2)]
apply (simp)
- unfolding setsum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
+ unfolding sum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
apply simp
done
finally show ?thesis
@@ -2088,16 +2088,16 @@
unfolding fps_eq_iff by blast
qed
-lemma fps_divide_X_minus1_setsum:
- "a /((1::'a::field fps) - X) = Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
+lemma fps_divide_X_minus1_sum:
+ "a /((1::'a::field fps) - X) = Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
proof -
let ?X = "1 - (X::'a fps)"
have th0: "?X $ 0 \<noteq> 0"
by simp
- have "a /?X = ?X * Abs_fps (\<lambda>n::nat. setsum (op $ a) {0..n}) * inverse ?X"
- using fps_divide_X_minus1_setsum_lemma[of a, symmetric] th0
+ have "a /?X = ?X * Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) * inverse ?X"
+ using fps_divide_X_minus1_sum_lemma[of a, symmetric] th0
by (simp add: fps_divide_def mult.assoc)
- also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n::nat. setsum (op $ a) {0..n}) "
+ also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n::nat. sum (op $ a) {0..n}) "
by (simp add: ac_simps)
finally show ?thesis
by (simp add: inverse_mult_eq_1[OF th0])
@@ -2228,12 +2228,12 @@
by auto
have d: "({0..k} - {i}) \<inter> {i} = {}"
using i by auto
- from H have "n = setsum (nth xs) {0..k}"
+ from H have "n = sum (nth xs) {0..k}"
apply (simp add: natpermute_def)
- apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_setsum_nth)
+ apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth)
done
- also have "\<dots> = n + setsum (nth xs) ({0..k} - {i})"
- unfolding setsum.union_disjoint[OF f d, unfolded eqs] using i by simp
+ also have "\<dots> = n + sum (nth xs) ({0..k} - {i})"
+ unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp
finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
by auto
from H have xsl: "length xs = k+1"
@@ -2265,11 +2265,11 @@
done
have xsl: "length xs = k + 1"
by (simp only: xs length_replicate length_list_update)
- have "sum_list xs = setsum (nth xs) {0..<k+1}"
- unfolding sum_list_setsum_nth xsl ..
- also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
- by (rule setsum.cong) (simp_all add: xs del: replicate.simps)
- also have "\<dots> = n" using i by (simp add: setsum.delta)
+ have "sum_list xs = sum (nth xs) {0..<k+1}"
+ unfolding sum_list_sum_nth xsl ..
+ also have "\<dots> = sum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
+ by (rule sum.cong) (simp_all add: xs del: replicate.simps)
+ also have "\<dots> = n" using i by (simp add: sum.delta)
finally have "xs \<in> natpermute n (k + 1)"
using xsl unfolding natpermute_def mem_Collect_eq by blast
then show "xs \<in> ?A"
@@ -2282,7 +2282,7 @@
fixes m :: nat
and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
shows "(setprod a {0 .. m}) $ n =
- setsum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
+ sum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
(is "?P m n")
proof (induct m arbitrary: n rule: nat_less_induct)
fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
@@ -2309,7 +2309,7 @@
apply (simp add: Suc)
unfolding natpermute_split[of m "m + 1", simplified, of n,
unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
- apply (subst setsum.UNION_disjoint)
+ apply (subst sum.UNION_disjoint)
apply simp
apply simp
unfolding image_Collect[symmetric]
@@ -2318,11 +2318,11 @@
apply (rule natpermute_finite)
apply (clarsimp simp add: set_eq_iff)
apply auto
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (rule refl)
- unfolding setsum_distrib_right
+ unfolding sum_distrib_right
apply (rule sym)
- apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in setsum.reindex_cong)
+ apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in sum.reindex_cong)
apply (simp add: inj_on_def)
apply auto
unfolding setprod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
@@ -2336,7 +2336,7 @@
lemma fps_power_nth_Suc:
fixes m :: nat
and a :: "'a::comm_ring_1 fps"
- shows "(a ^ Suc m)$n = setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
+ shows "(a ^ Suc m)$n = sum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
proof -
have th0: "a^Suc m = setprod (\<lambda>i. a) {0..m}"
by (simp add: setprod_constant)
@@ -2347,7 +2347,7 @@
fixes m :: nat
and a :: "'a::comm_ring_1 fps"
shows "(a ^m)$n =
- (if m=0 then 1$n else setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
+ (if m=0 then 1$n else sum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc)
lemma fps_nth_power_0:
@@ -2423,13 +2423,13 @@
from v have "k = sum_list v" by (simp add: A_def natpermute_def)
also have "\<dots> = (\<Sum>i=0..m. v ! i)"
- by (simp add: sum_list_setsum_nth atLeastLessThanSuc_atLeastAtMost del: setsum_op_ivl_Suc)
+ by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum_op_ivl_Suc)
also from j have "{0..m} = insert j ({0..m}-{j})" by auto
also from j have "(\<Sum>i\<in>\<dots>. v ! i) = k + (\<Sum>i\<in>{0..m}-{j}. v ! i)"
- by (subst setsum.insert) simp_all
+ by (subst sum.insert) simp_all
finally have "(\<Sum>i\<in>{0..m}-{j}. v ! i) = 0" by simp
hence zero: "v ! i = 0" if "i \<in> {0..m}-{j}" for i using that
- by (subst (asm) setsum_eq_0_iff) auto
+ by (subst (asm) sum_eq_0_iff) auto
from j have "{0..m} = insert j ({0..m} - {j})" by auto
also from j have "(\<Prod>i\<in>\<dots>. f $ (v ! i)) = f $ k * (\<Prod>i\<in>{0..m} - {j}. f $ (v ! i))"
@@ -2445,7 +2445,7 @@
also have "natpermute k (m+1) = A \<union> B" unfolding A_def B_def by blast
also have "(\<Sum>v\<in>\<dots>. \<Prod>j = 0..m. f $ (v ! j)) =
(\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) + (\<Sum>v\<in>B. \<Prod>j = 0..m. f $ (v ! j))"
- by (intro setsum.union_disjoint) simp_all
+ by (intro sum.union_disjoint) simp_all
also have "(\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)"
by (simp add: A card_A)
finally show ?thesis by (simp add: B_def)
@@ -2472,7 +2472,7 @@
using that elem_le_sum_list_nat[of i v] unfolding natpermute_def
by (auto simp: set_conv_nth dest!: spec[of _ i])
hence "?h f = ?h g"
- by (intro setsum.cong refl setprod.cong less lessI) (auto simp: natpermute_def)
+ by (intro sum.cong refl setprod.cong less lessI) (auto simp: natpermute_def)
finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)"
by simp
with assms show "f $ k = g $ k"
@@ -2544,16 +2544,16 @@
have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using Suc by auto
have d: "{0 .. n1} \<inter> {n} = {}" using Suc by auto
have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
- apply (rule setsum.cong)
+ apply (rule sum.cong)
using H Suc
apply auto
done
have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
- unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq] seq
+ unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq
using startsby_zero_power_nth_same[OF a0]
by simp
have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
- unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq]
+ unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq]
using startsby_zero_power_nth_same[OF a0]
by simp
from \<open>?lhs\<close>[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
@@ -2575,7 +2575,7 @@
| "radical r 0 a (Suc n) = 0"
| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
| "radical r (Suc k) a (Suc n) =
- (a$ Suc n - setsum (\<lambda>xs. setprod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
+ (a$ Suc n - sum (\<lambda>xs. setprod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
{xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
(of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
by pat_completeness auto
@@ -2601,11 +2601,11 @@
using i by auto
from xs have "Suc n = sum_list xs"
by (simp add: natpermute_def)
- also have "\<dots> = setsum (nth xs) {0..<Suc k}" using xs
- by (simp add: natpermute_def sum_list_setsum_nth)
- also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
- unfolding eqs setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
- unfolding setsum.union_disjoint[OF fths(2) fths(3) d(2)]
+ also have "\<dots> = sum (nth xs) {0..<Suc k}" using xs
+ by (simp add: natpermute_def sum_list_sum_nth)
+ also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
+ unfolding eqs sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+ unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
by simp
finally show ?thesis using c' by simp
qed
@@ -2684,8 +2684,8 @@
using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
by (metis natpermute_finite)+
let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
- have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
- proof (rule setsum.cong)
+ have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+ proof (rule sum.cong)
fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
@@ -2701,13 +2701,13 @@
using i r0 by (simp add: setprod_gen_delta)
finally show ?ths .
qed rule
- then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+ then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
by (simp add: natpermute_max_card[OF \<open>n \<noteq> 0\<close>, simplified])
- also have "\<dots> = a$n - setsum ?f ?Pnknn"
+ also have "\<dots> = a$n - sum ?f ?Pnknn"
unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
- finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
- have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
- unfolding fps_power_nth_Suc setsum.union_disjoint[OF f d, unfolded eq] ..
+ finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
+ have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
+ unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
also have "\<dots> = a$n" unfolding fn by simp
finally show ?thesis .
qed
@@ -2750,8 +2750,8 @@
using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
by (metis natpermute_finite)+
let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
- have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
- proof(rule setsum.cong2)
+ have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
+ proof(rule sum.cong2)
fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
@@ -2763,13 +2763,13 @@
unfolding setprod_gen_delta[OF fK] using i r0 by simp
finally show ?ths .
qed
- then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
+ then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
by (simp add: natpermute_max_card[OF nz, simplified])
- also have "\<dots> = a$n - setsum ?f ?Pnknn"
+ also have "\<dots> = a$n - sum ?f ?Pnknn"
unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
- finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
- have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
- unfolding fps_power_nth_Suc setsum.union_disjoint[OF f d, unfolded eq] ..
+ finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
+ have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
+ unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
also have "\<dots> = a$n" unfolding fn by simp
finally have "?r ^ Suc k $ n = a $n" .}
ultimately show "?r ^ Suc k $ n = a $n" by (cases n, auto)
@@ -2819,8 +2819,8 @@
by (metis natpermute_finite)+
let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
- have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
- proof (rule setsum.cong)
+ have "sum ?g ?Pnkn = sum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
+ proof (rule sum.cong)
fix v
assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
@@ -2836,10 +2836,10 @@
using i by (simp add: setprod_gen_delta)
finally show ?ths .
qed rule
- then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
+ then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
by (simp add: natpermute_max_card[OF nz, simplified])
- have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
- proof (rule setsum.cong, rule refl, rule setprod.cong, simp)
+ have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn"
+ proof (rule sum.cong, rule refl, rule setprod.cong, simp)
fix xs i
assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
have False if c: "n \<le> xs ! i"
@@ -2855,11 +2855,11 @@
using i by auto
from xs have "n = sum_list xs"
by (simp add: natpermute_def)
- also have "\<dots> = setsum (nth xs) {0..<Suc k}"
- using xs by (simp add: natpermute_def sum_list_setsum_nth)
- also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
- unfolding eqs setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
- unfolding setsum.union_disjoint[OF fths(2) fths(3) d(2)]
+ also have "\<dots> = sum (nth xs) {0..<Suc k}"
+ using xs by (simp add: natpermute_def sum_list_sum_nth)
+ also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
+ unfolding eqs sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
+ unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
by simp
finally show ?thesis using c' by simp
qed
@@ -2870,15 +2870,15 @@
by (simp add: field_simps del: of_nat_Suc)
from \<open>?lhs\<close> have "b$n = a^Suc k $ n"
by (simp add: fps_eq_iff)
- also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
+ also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn"
unfolding fps_power_nth_Suc
- using setsum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
+ using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
unfolded eq, of ?g] by simp
- also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn"
+ also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn"
unfolding th0 th1 ..
- finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn"
+ finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn"
by simp
- then have "a$n = (b$n - setsum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
+ then have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
apply -
apply (rule eq_divide_imp')
using r00
@@ -3081,31 +3081,31 @@
proof -
have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n
proof -
- have "(fps_deriv (a oo b))$n = setsum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
- by (simp add: fps_compose_def field_simps setsum_distrib_left del: of_nat_Suc)
- also have "\<dots> = setsum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
+ have "(fps_deriv (a oo b))$n = sum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
+ by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc)
+ also have "\<dots> = sum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
- also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
+ also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
unfolding fps_mult_left_const_nth by (simp add: field_simps)
- also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
+ also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
unfolding fps_mult_nth ..
- also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
- apply (rule setsum.mono_neutral_right)
- apply (auto simp add: mult_delta_left setsum.delta not_le)
+ also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
+ apply (rule sum.mono_neutral_right)
+ apply (auto simp add: mult_delta_left sum.delta not_le)
done
- also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
+ also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
unfolding fps_deriv_nth
- by (rule setsum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
+ by (rule sum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
finally have th0: "(fps_deriv (a oo b))$n =
- setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
-
- have "(((fps_deriv a) oo b) * (fps_deriv b))$n = setsum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
+ sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
+
+ have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
unfolding fps_mult_nth by (simp add: ac_simps)
- also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
- unfolding fps_deriv_nth fps_compose_nth setsum_distrib_left mult.assoc
- apply (rule setsum.cong)
+ also have "\<dots> = sum (\<lambda>i. sum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
+ unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc
+ apply (rule sum.cong)
apply (rule refl)
- apply (rule setsum.mono_neutral_left)
+ apply (rule sum.mono_neutral_left)
apply (simp_all add: subset_eq)
apply clarify
apply (subgoal_tac "b^i$x = 0")
@@ -3113,10 +3113,10 @@
apply (rule startsby_zero_power_prefix[OF b0, rule_format])
apply simp
done
- also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
- unfolding setsum_distrib_left
- apply (subst setsum.commute)
- apply (rule setsum.cong, rule refl)+
+ also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
+ unfolding sum_distrib_left
+ apply (subst sum.commute)
+ apply (rule sum.cong, rule refl)+
apply simp
done
finally show ?thesis
@@ -3133,10 +3133,10 @@
by (simp add: fps_mult_nth)
next
case (Suc m)
- have "((1 + X)*a) $ n = setsum (\<lambda>i. (1 + X) $ i * a $ (n - i)) {0..n}"
+ have "((1 + X)*a) $ n = sum (\<lambda>i. (1 + X) $ i * a $ (n - i)) {0..n}"
by (simp add: fps_mult_nth)
- also have "\<dots> = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
- unfolding Suc by (rule setsum.mono_neutral_right) auto
+ also have "\<dots> = sum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
+ unfolding Suc by (rule sum.mono_neutral_right) auto
also have "\<dots> = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
by (simp add: Suc)
finally show ?thesis .
@@ -3147,11 +3147,11 @@
lemma fps_poly_sum_X:
assumes "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
- shows "a = setsum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
+ shows "a = sum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
proof -
have "a$i = ?r$i" for i
- unfolding fps_setsum_nth fps_mult_left_const_nth X_power_nth
- by (simp add: mult_delta_right setsum.delta' assms)
+ unfolding fps_sum_nth fps_mult_left_const_nth X_power_nth
+ by (simp add: mult_delta_right sum.delta' assms)
then show ?thesis
unfolding fps_eq_iff by blast
qed
@@ -3163,7 +3163,7 @@
where
"compinv a 0 = X$0"
| "compinv a (Suc n) =
- (X$ Suc n - setsum (\<lambda>i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
+ (X$ Suc n - sum (\<lambda>i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
definition "fps_inv a = Abs_fps (compinv a)"
@@ -3184,10 +3184,10 @@
by (simp add: fps_compose_nth fps_inv_def)
next
case (Suc n1)
- have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
+ have "?i $ n = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
- also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
- (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
+ also have "\<dots> = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
+ (X$ Suc n1 - sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
using a0 a1 Suc by (simp add: fps_inv_def)
also have "\<dots> = X$n" using Suc by simp
finally show ?thesis .
@@ -3202,7 +3202,7 @@
where
"gcompinv b a 0 = b$0"
| "gcompinv b a (Suc n) =
- (b$ Suc n - setsum (\<lambda>i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
+ (b$ Suc n - sum (\<lambda>i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
definition "fps_ginv b a = Abs_fps (gcompinv b a)"
@@ -3223,10 +3223,10 @@
by (simp add: fps_compose_nth fps_ginv_def)
next
case (Suc n1)
- have "?i $ n = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
+ have "?i $ n = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
- also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
- (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
+ also have "\<dots> = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
+ (b$ Suc n1 - sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
using a0 a1 Suc by (simp add: fps_ginv_def)
also have "\<dots> = b$n" using Suc by simp
finally show ?thesis .
@@ -3246,30 +3246,30 @@
done
lemma fps_compose_1[simp]: "1 oo a = 1"
- by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
+ by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
lemma fps_compose_0[simp]: "0 oo a = 0"
by (simp add: fps_eq_iff fps_compose_nth)
lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)"
- by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum.neutral)
+ by (auto simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral)
lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
- by (simp add: fps_eq_iff fps_compose_nth field_simps setsum.distrib)
-
-lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
+ by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib)
+
+lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\<lambda>i. f i oo a) S"
proof (cases "finite S")
case True
show ?thesis
proof (rule finite_induct[OF True])
- show "setsum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
+ show "sum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
by simp
next
fix x F
assume fF: "finite F"
and xF: "x \<notin> F"
- and h: "setsum f F oo a = setsum (\<lambda>i. f i oo a) F"
- show "setsum f (insert x F) oo a = setsum (\<lambda>i. f i oo a) (insert x F)"
+ and h: "sum f F oo a = sum (\<lambda>i. f i oo a) F"
+ show "sum f (insert x F) oo a = sum (\<lambda>i. f i oo a) (insert x F)"
using fF xF h by (simp add: fps_compose_add_distrib)
qed
next
@@ -3278,15 +3278,15 @@
qed
lemma convolution_eq:
- "setsum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
- setsum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
- by (rule setsum.reindex_bij_witness[where i=fst and j="\<lambda>i. (i, n - i)"]) auto
+ "sum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
+ sum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
+ by (rule sum.reindex_bij_witness[where i=fst and j="\<lambda>i. (i, n - i)"]) auto
lemma product_composition_lemma:
assumes c0: "c$0 = (0::'a::idom)"
and d0: "d$0 = 0"
shows "((a oo c) * (b oo d))$n =
- setsum (\<lambda>(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}" (is "?l = ?r")
+ sum (\<lambda>(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}" (is "?l = ?r")
proof -
let ?S = "{(k::nat, m::nat). k + m \<le> n}"
have s: "?S \<subseteq> {0..n} \<times> {0..n}" by (auto simp add: subset_eq)
@@ -3294,18 +3294,18 @@
apply (rule finite_subset[OF s])
apply auto
done
- have "?r = setsum (\<lambda>i. setsum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
- apply (simp add: fps_mult_nth setsum_distrib_left)
- apply (subst setsum.commute)
- apply (rule setsum.cong)
+ have "?r = sum (\<lambda>i. sum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
+ apply (simp add: fps_mult_nth sum_distrib_left)
+ apply (subst sum.commute)
+ apply (rule sum.cong)
apply (auto simp add: field_simps)
done
also have "\<dots> = ?l"
- apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
- apply (rule setsum.cong)
+ apply (simp add: fps_mult_nth fps_compose_nth sum_product)
+ apply (rule sum.cong)
apply (rule refl)
- apply (simp add: setsum.cartesian_product mult.assoc)
- apply (rule setsum.mono_neutral_right[OF f])
+ apply (simp add: sum.cartesian_product mult.assoc)
+ apply (rule sum.mono_neutral_right[OF f])
apply (simp add: subset_eq)
apply presburger
apply clarsimp
@@ -3326,10 +3326,10 @@
assumes c0: "c$0 = (0::'a::idom)"
and d0: "d$0 = 0"
shows "((a oo c) * (b oo d))$n =
- setsum (\<lambda>k. setsum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}" (is "?l = ?r")
+ sum (\<lambda>k. sum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}" (is "?l = ?r")
unfolding product_composition_lemma[OF c0 d0]
- unfolding setsum.cartesian_product
- apply (rule setsum.mono_neutral_left)
+ unfolding sum.cartesian_product
+ apply (rule sum.mono_neutral_left)
apply simp
apply (clarsimp simp add: subset_eq)
apply clarsimp
@@ -3337,7 +3337,7 @@
apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
apply simp
unfolding fps_mult_nth
- apply (rule setsum.neutral)
+ apply (rule sum.neutral)
apply (clarsimp simp add: not_le)
apply (case_tac "x < aa")
apply (rule startsby_zero_power_prefix[OF c0, rule_format])
@@ -3349,9 +3349,9 @@
done
-lemma setsum_pair_less_iff:
- "setsum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
- setsum (\<lambda>s. setsum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
+lemma sum_pair_less_iff:
+ "sum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
+ sum (\<lambda>s. sum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
(is "?l = ?r")
proof -
let ?KM = "{(k,m). k + m \<le> n}"
@@ -3360,24 +3360,24 @@
by auto
show "?l = ?r "
unfolding th0
- apply (subst setsum.UNION_disjoint)
+ apply (subst sum.UNION_disjoint)
apply auto
- apply (subst setsum.UNION_disjoint)
+ apply (subst sum.UNION_disjoint)
apply auto
done
qed
lemma fps_compose_mult_distrib_lemma:
assumes c0: "c$0 = (0::'a::idom)"
- shows "((a oo c) * (b oo c))$n = setsum (\<lambda>s. setsum (\<lambda>i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
+ shows "((a oo c) * (b oo c))$n = sum (\<lambda>s. sum (\<lambda>i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
- unfolding setsum_pair_less_iff[where a = "\<lambda>k. a$k" and b="\<lambda>m. b$m" and c="\<lambda>s. (c ^ s)$n" and n = n] ..
+ unfolding sum_pair_less_iff[where a = "\<lambda>k. a$k" and b="\<lambda>m. b$m" and c="\<lambda>s. (c ^ s)$n" and n = n] ..
lemma fps_compose_mult_distrib:
assumes c0: "c $ 0 = (0::'a::idom)"
shows "(a * b) oo c = (a oo c) * (b oo c)"
apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
- apply (simp add: fps_compose_nth fps_mult_nth setsum_distrib_right)
+ apply (simp add: fps_compose_nth fps_mult_nth sum_distrib_right)
done
lemma fps_compose_setprod_distrib:
@@ -3420,13 +3420,13 @@
qed
lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
- by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
+ by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric])
lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
- by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
+ by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
lemma fps_inverse_compose:
assumes b0: "(b$0 :: 'a::field) = 0"
@@ -3498,7 +3498,7 @@
qed
lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
- by (simp add: fps_eq_iff fps_compose_nth setsum_distrib_left mult.assoc)
+ by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc)
lemma fps_const_mult_apply_right:
"(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
@@ -3511,16 +3511,16 @@
proof -
have "?l$n = ?r$n" for n
proof -
- have "?l$n = (setsum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
+ have "?l$n = (sum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
- setsum_distrib_left mult.assoc fps_setsum_nth)
- also have "\<dots> = ((setsum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
- by (simp add: fps_compose_setsum_distrib)
+ sum_distrib_left mult.assoc fps_sum_nth)
+ also have "\<dots> = ((sum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
+ by (simp add: fps_compose_sum_distrib)
also have "\<dots> = ?r$n"
- apply (simp add: fps_compose_nth fps_setsum_nth setsum_distrib_right mult.assoc)
- apply (rule setsum.cong)
+ apply (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc)
+ apply (rule sum.cong)
apply (rule refl)
- apply (rule setsum.mono_neutral_right)
+ apply (rule sum.mono_neutral_right)
apply (auto simp add: not_le)
apply (erule startsby_zero_power_prefix[OF b0, rule_format])
done
@@ -3552,7 +3552,7 @@
next
case 2
then show ?thesis
- by (simp add: fps_compose_nth mult_delta_left setsum.delta)
+ by (simp add: fps_compose_nth mult_delta_left sum.delta)
qed
qed
then show ?thesis
@@ -3688,7 +3688,7 @@
lemma fps_compose_linear:
"fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * X) = Abs_fps (\<lambda>n. c^n * f $ n)"
by (simp add: fps_eq_iff fps_compose_def power_mult_distrib
- if_distrib setsum.delta' cong: if_cong)
+ if_distrib sum.delta' cong: if_cong)
subsection \<open>Elementary series\<close>
@@ -3796,7 +3796,7 @@
lemma Ec_E1_eq: "E (1::'a::field_char_0) oo (fps_const c * X) = E c"
apply (auto simp add: fps_eq_iff E_def fps_compose_def power_mult_distrib)
- apply (simp add: cond_value_iff cond_application_beta setsum.delta' cong del: if_weak_cong)
+ apply (simp add: cond_value_iff cond_application_beta sum.delta' cong del: if_weak_cong)
done
@@ -4066,7 +4066,7 @@
text \<open>Vandermonde's Identity as a consequence.\<close>
lemma gbinomial_Vandermonde:
- "setsum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
+ "sum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
proof -
let ?ba = "fps_binomial a"
let ?bb = "fps_binomial b"
@@ -4076,23 +4076,23 @@
qed
lemma binomial_Vandermonde:
- "setsum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
+ "sum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n]
by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric]
- of_nat_setsum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
-
-lemma binomial_Vandermonde_same: "setsum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
+ of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
+
+lemma binomial_Vandermonde_same: "sum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
using binomial_Vandermonde[of n n n, symmetric]
unfolding mult_2
apply (simp add: power2_eq_square)
- apply (rule setsum.cong)
+ apply (rule sum.cong)
apply (auto intro: binomial_symmetric)
done
lemma Vandermonde_pochhammer_lemma:
fixes a :: "'a::field_char_0"
assumes b: "\<forall>j\<in>{0 ..<n}. b \<noteq> of_nat j"
- shows "setsum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
+ shows "sum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
(of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} =
pochhammer (- (a + b)) n / pochhammer (- b) n"
(is "?l = ?r")
@@ -4224,7 +4224,7 @@
apply (simp add: th00)
unfolding gbinomial_pochhammer
using bn0
- apply (simp add: setsum_distrib_right setsum_distrib_left field_simps)
+ apply (simp add: sum_distrib_right sum_distrib_left field_simps)
done
finally show ?thesis by simp
qed
@@ -4232,7 +4232,7 @@
lemma Vandermonde_pochhammer:
fixes a :: "'a::field_char_0"
assumes c: "\<forall>i \<in> {0..< n}. c \<noteq> - of_nat i"
- shows "setsum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
+ shows "sum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
(of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
proof -
let ?a = "- a"
@@ -4253,7 +4253,7 @@
by (simp add: pochhammer_eq_0_iff)
from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
show ?thesis
- using nz by (simp add: field_simps setsum_distrib_left)
+ using nz by (simp add: field_simps sum_distrib_left)
qed
@@ -4558,7 +4558,7 @@
have th0: "(fps_const c * X) $ 0 = 0" by simp
show ?thesis unfolding gp[OF th0, symmetric]
by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
- fps_compose_nth power_mult_distrib cond_value_iff setsum.delta' cong del: if_weak_cong)
+ fps_compose_nth power_mult_distrib cond_value_iff sum.delta' cong del: if_weak_cong)
qed
lemma F_B[simp]: "F [-a] [] (- 1) = fps_binomial a"
@@ -4620,10 +4620,10 @@
else 0)"
by (auto simp add: pochhammer_eq_0_iff)
-lemma setsum_eq_if: "setsum f {(n::nat) .. m} = (if m < n then 0 else f n + setsum f {n+1 .. m})"
+lemma sum_eq_if: "sum f {(n::nat) .. m} = (if m < n then 0 else f n + sum f {n+1 .. m})"
apply simp
- apply (subst setsum.insert[symmetric])
- apply (auto simp add: not_less setsum_head_Suc)
+ apply (subst sum.insert[symmetric])
+ apply (auto simp add: not_less sum_head_Suc)
done
lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))"
--- a/src/HOL/Library/Groups_Big_Fun.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Groups_Big_Fun.thy Mon Oct 17 14:37:32 2016 +0200
@@ -216,11 +216,11 @@
sublocale Sum_any: comm_monoid_fun plus 0
defines Sum_any = Sum_any.G
- rewrites "comm_monoid_set.F plus 0 = setsum"
+ rewrites "comm_monoid_set.F plus 0 = sum"
proof -
show "comm_monoid_fun plus 0" ..
then interpret Sum_any: comm_monoid_fun plus 0 .
- from setsum_def show "comm_monoid_set.F plus 0 = setsum" by (auto intro: sym)
+ from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym)
qed
end
@@ -240,7 +240,7 @@
note assms
moreover have "{a. g a * r \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
ultimately show ?thesis
- by (simp add: setsum_distrib_right Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
+ by (simp add: sum_distrib_right Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
qed
lemma Sum_any_right_distrib:
@@ -251,7 +251,7 @@
note assms
moreover have "{a. r * g a \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
ultimately show ?thesis
- by (simp add: setsum_distrib_left Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
+ by (simp add: sum_distrib_left Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
qed
lemma Sum_any_product:
@@ -266,7 +266,7 @@
ultimately show ?thesis using assms
by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g]
Sum_any.expand_superset [of "{a. f a \<noteq> 0}"] Sum_any.expand_superset [of "{b. g b \<noteq> 0}"]
- setsum_product)
+ sum_product)
qed
lemma Sum_any_eq_zero_iff [simp]:
@@ -325,7 +325,7 @@
have "{a. c ^ f a \<noteq> 1} \<subseteq> {a. f a \<noteq> 0}"
by (auto intro: ccontr)
with assms show ?thesis
- by (simp add: Sum_any.expand_set Prod_any.expand_superset power_setsum)
+ by (simp add: Sum_any.expand_set Prod_any.expand_superset power_sum)
qed
end
--- a/src/HOL/Library/Indicator_Function.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Indicator_Function.thy Mon Oct 17 14:37:32 2016 +0200
@@ -77,21 +77,21 @@
lemma (* FIXME unnamed!? *)
fixes f :: "'a \<Rightarrow> 'b::semiring_1"
assumes "finite A"
- shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator B x) = (\<Sum>x \<in> A \<inter> B. f x)"
- and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator B x * f x) = (\<Sum>x \<in> A \<inter> B. f x)"
+ shows sum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator B x) = (\<Sum>x \<in> A \<inter> B. f x)"
+ and sum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator B x * f x) = (\<Sum>x \<in> A \<inter> B. f x)"
unfolding indicator_def
- using assms by (auto intro!: setsum.mono_neutral_cong_right split: if_split_asm)
+ using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm)
-lemma setsum_indicator_eq_card:
+lemma sum_indicator_eq_card:
assumes "finite A"
shows "(\<Sum>x \<in> A. indicator B x) = card (A Int B)"
- using setsum_mult_indicator [OF assms, of "\<lambda>x. 1::nat"]
- unfolding card_eq_setsum by simp
+ using sum_mult_indicator [OF assms, of "\<lambda>x. 1::nat"]
+ unfolding card_eq_sum by simp
-lemma setsum_indicator_scaleR[simp]:
+lemma sum_indicator_scaleR[simp]:
"finite A \<Longrightarrow>
(\<Sum>x \<in> A. indicator (B x) (g x) *\<^sub>R f x) = (\<Sum>x \<in> {x\<in>A. g x \<in> B x}. f x :: 'a::real_vector)"
- by (auto intro!: setsum.mono_neutral_cong_right split: if_split_asm simp: indicator_def)
+ by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm simp: indicator_def)
lemma LIMSEQ_indicator_incseq:
assumes "incseq A"
--- a/src/HOL/Library/Multiset.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Multiset.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1330,7 +1330,7 @@
unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def)
definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where
- "size_multiset f M = setsum (wcount f M) (set_mset M)"
+ "size_multiset f M = sum (wcount f M) (set_mset M)"
lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def]
@@ -1358,16 +1358,16 @@
lemma size_single: "size {#b#} = 1"
by (simp add: size_multiset_overloaded_def size_multiset_single)
-lemma setsum_wcount_Int:
- "finite A \<Longrightarrow> setsum (wcount f N) (A \<inter> set_mset N) = setsum (wcount f N) A"
+lemma sum_wcount_Int:
+ "finite A \<Longrightarrow> sum (wcount f N) (A \<inter> set_mset N) = sum (wcount f N) A"
by (induct rule: finite_induct)
(simp_all add: Int_insert_left wcount_def count_eq_zero_iff)
lemma size_multiset_union [simp]:
"size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N"
-apply (simp add: size_multiset_def setsum_Un_nat setsum.distrib setsum_wcount_Int wcount_union)
+apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union)
apply (subst Int_commute)
-apply (simp add: setsum_wcount_Int)
+apply (simp add: sum_wcount_Int)
done
lemma size_multiset_add_mset [simp]:
@@ -1392,7 +1392,7 @@
lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M"
apply (unfold size_multiset_overloaded_eq)
-apply (drule setsum_SucD)
+apply (drule sum_SucD)
apply auto
done
@@ -1663,7 +1663,7 @@
moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y
by simp
ultimately show ?case
- by (auto simp: setsum.distrib setsum.delta' intro!: setsum.mono_neutral_left)
+ by (auto simp: sum.distrib sum.delta' intro!: sum.mono_neutral_left)
qed
lemma image_mset_subseteq_mono: "A \<subseteq># B \<Longrightarrow> image_mset f A \<subseteq># image_mset f B"
@@ -2154,8 +2154,8 @@
"sum_mset (replicate_mset n a) = of_nat n * a"
by (induct n) (simp_all add: algebra_simps)
-lemma setsum_unfold_sum_mset:
- "setsum f A = sum_mset (image_mset f (mset_set A))"
+lemma sum_unfold_sum_mset:
+ "sum f A = sum_mset (image_mset f (mset_set A))"
by (cases "finite A") (induct A rule: finite_induct, simp_all)
lemma sum_mset_delta: "sum_mset (image_mset (\<lambda>x. if x = y then c else 0) A) = c * count A y"
@@ -2186,12 +2186,12 @@
next
case (add x M) then show ?case
by (cases "x \<in> set_mset M")
- (simp_all add: size_multiset_overloaded_eq not_in_iff setsum.If_cases Diff_eq[symmetric]
- setsum.remove)
+ (simp_all add: size_multiset_overloaded_eq not_in_iff sum.If_cases Diff_eq[symmetric]
+ sum.remove)
qed
lemma size_mset_set [simp]: "size (mset_set A) = card A"
-by (simp only: size_eq_sum_mset card_eq_setsum setsum_unfold_sum_mset)
+by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset)
syntax (ASCII)
"_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10)
@@ -2235,13 +2235,13 @@
lemma in_Union_mset_iff[iff]: "x \<in># \<Union># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)"
by (induct MM) auto
-lemma count_setsum:
- "count (setsum f A) x = setsum (\<lambda>a. count (f a) x) A"
+lemma count_sum:
+ "count (sum f A) x = sum (\<lambda>a. count (f a) x) A"
by (induct A rule: infinite_finite_induct) simp_all
-lemma setsum_eq_empty_iff:
+lemma sum_eq_empty_iff:
assumes "finite A"
- shows "setsum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})"
+ shows "sum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})"
using assms by induct simp_all
lemma Union_mset_empty_conv[simp]: "\<Union># M = {#} \<longleftrightarrow> (\<forall>i\<in>#M. i = {#})"
--- a/src/HOL/Library/Multiset_Permutations.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Multiset_Permutations.thy Mon Oct 17 14:37:32 2016 +0200
@@ -192,9 +192,9 @@
also have "\<dots> * (\<Prod>x\<in>set_mset A. fact (count A x)) =
(\<Sum>x\<in>set_mset A. card (permutations_of_multiset (A - {#x#})) *
(\<Prod>y\<in>set_mset A. fact (count A y)))"
- by (subst setsum_distrib_right) simp_all
+ by (subst sum_distrib_right) simp_all
also have "\<dots> = (\<Sum>x\<in>set_mset A. count A x * fact (size A - 1))"
- proof (intro setsum.cong refl)
+ proof (intro sum.cong refl)
fix x assume x: "x \<in># A"
have "card (permutations_of_multiset (A - {#x#})) * (\<Prod>y\<in>set_mset A. fact (count A y)) =
count A x * (card (permutations_of_multiset (A - {#x#})) *
@@ -206,7 +206,7 @@
qed
also have "(\<Sum>x\<in>set_mset A. count A x * fact (size A - 1)) =
size A * fact (size A - 1)"
- by (simp add: setsum_distrib_right size_multiset_overloaded_eq)
+ by (simp add: sum_distrib_right size_multiset_overloaded_eq)
also from remove.hyps have "\<dots> = fact (size A)"
by (cases "size A") auto
finally show ?case .
@@ -507,4 +507,4 @@
value [code] "permutations_of_set (set ''abcd'')"
-end
\ No newline at end of file
+end
--- a/src/HOL/Library/Nat_Bijection.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Nat_Bijection.thy Mon Oct 17 14:37:32 2016 +0200
@@ -272,7 +272,7 @@
subsubsection \<open>From sets to naturals\<close>
definition set_encode :: "nat set \<Rightarrow> nat"
- where "set_encode = setsum (op ^ 2)"
+ where "set_encode = sum (op ^ 2)"
lemma set_encode_empty [simp]: "set_encode {} = 0"
by (simp add: set_encode_def)
@@ -381,7 +381,7 @@
"n = set_encode B" "finite B"
by (metis finite_set_decode set_decode_inverse)
with assms show ?thesis
- by auto (simp add: set_encode_def add.commute setsum.subset_diff)
+ by auto (simp add: set_encode_def add.commute sum.subset_diff)
qed
then show ?thesis
by (metis le_add1)
--- a/src/HOL/Library/Permutations.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Permutations.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1191,8 +1191,8 @@
lemma permutes_in_seg: "p permutes {1 ..n} \<Longrightarrow> i \<in> {1..n} \<Longrightarrow> 1 \<le> p i \<and> p i \<le> n"
by (simp add: permutes_def) metis
-lemma setsum_permutations_inverse:
- "setsum f {p. p permutes S} = setsum (\<lambda>p. f(inv p)) {p. p permutes S}"
+lemma sum_permutations_inverse:
+ "sum f {p. p permutes S} = sum (\<lambda>p. f(inv p)) {p. p permutes S}"
(is "?lhs = ?rhs")
proof -
let ?S = "{p . p permutes S}"
@@ -1209,18 +1209,18 @@
qed
have th1: "inv ` ?S = ?S"
using image_inverse_permutations by blast
- have th2: "?rhs = setsum (f \<circ> inv) ?S"
+ have th2: "?rhs = sum (f \<circ> inv) ?S"
by (simp add: o_def)
- from setsum.reindex[OF th0, of f] show ?thesis unfolding th1 th2 .
+ from sum.reindex[OF th0, of f] show ?thesis unfolding th1 th2 .
qed
lemma setum_permutations_compose_left:
assumes q: "q permutes S"
- shows "setsum f {p. p permutes S} = setsum (\<lambda>p. f(q \<circ> p)) {p. p permutes S}"
+ shows "sum f {p. p permutes S} = sum (\<lambda>p. f(q \<circ> p)) {p. p permutes S}"
(is "?lhs = ?rhs")
proof -
let ?S = "{p. p permutes S}"
- have th0: "?rhs = setsum (f \<circ> (op \<circ> q)) ?S"
+ have th0: "?rhs = sum (f \<circ> (op \<circ> q)) ?S"
by (simp add: o_def)
have th1: "inj_on (op \<circ> q) ?S"
proof (auto simp add: inj_on_def)
@@ -1235,16 +1235,16 @@
qed
have th3: "(op \<circ> q) ` ?S = ?S"
using image_compose_permutations_left[OF q] by auto
- from setsum.reindex[OF th1, of f] show ?thesis unfolding th0 th1 th3 .
+ from sum.reindex[OF th1, of f] show ?thesis unfolding th0 th1 th3 .
qed
lemma sum_permutations_compose_right:
assumes q: "q permutes S"
- shows "setsum f {p. p permutes S} = setsum (\<lambda>p. f(p \<circ> q)) {p. p permutes S}"
+ shows "sum f {p. p permutes S} = sum (\<lambda>p. f(p \<circ> q)) {p. p permutes S}"
(is "?lhs = ?rhs")
proof -
let ?S = "{p. p permutes S}"
- have th0: "?rhs = setsum (f \<circ> (\<lambda>p. p \<circ> q)) ?S"
+ have th0: "?rhs = sum (f \<circ> (\<lambda>p. p \<circ> q)) ?S"
by (simp add: o_def)
have th1: "inj_on (\<lambda>p. p \<circ> q) ?S"
proof (auto simp add: inj_on_def)
@@ -1259,18 +1259,18 @@
qed
have th3: "(\<lambda>p. p \<circ> q) ` ?S = ?S"
using image_compose_permutations_right[OF q] by auto
- from setsum.reindex[OF th1, of f]
+ from sum.reindex[OF th1, of f]
show ?thesis unfolding th0 th1 th3 .
qed
subsection \<open>Sum over a set of permutations (could generalize to iteration)\<close>
-lemma setsum_over_permutations_insert:
+lemma sum_over_permutations_insert:
assumes fS: "finite S"
and aS: "a \<notin> S"
- shows "setsum f {p. p permutes (insert a S)} =
- setsum (\<lambda>b. setsum (\<lambda>q. f (Fun.swap a b id \<circ> q)) {p. p permutes S}) (insert a S)"
+ shows "sum f {p. p permutes (insert a S)} =
+ sum (\<lambda>b. sum (\<lambda>q. f (Fun.swap a b id \<circ> q)) {p. p permutes S}) (insert a S)"
proof -
have th0: "\<And>f a b. (\<lambda>(b,p). f (Fun.swap a b id \<circ> p)) = f \<circ> (\<lambda>(b,p). Fun.swap a b id \<circ> p)"
by (simp add: fun_eq_iff)
@@ -1280,10 +1280,10 @@
by blast
show ?thesis
unfolding permutes_insert
- unfolding setsum.cartesian_product
+ unfolding sum.cartesian_product
unfolding th1[symmetric]
unfolding th0
- proof (rule setsum.reindex)
+ proof (rule sum.reindex)
let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"
let ?P = "{p. p permutes S}"
{
--- a/src/HOL/Library/Polynomial.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Polynomial.thy Mon Oct 17 14:37:32 2016 +0200
@@ -507,8 +507,8 @@
also note pCons.IH
also have "a + x * (\<Sum>i\<le>degree p. coeff p i * x ^ i) =
coeff ?p' 0 * x^0 + (\<Sum>i\<le>degree p. coeff ?p' (Suc i) * x^Suc i)"
- by (simp add: field_simps setsum_distrib_left coeff_pCons)
- also note setsum_atMost_Suc_shift[symmetric]
+ by (simp add: field_simps sum_distrib_left coeff_pCons)
+ also note sum_atMost_Suc_shift[symmetric]
also note degree_pCons_eq[OF \<open>p \<noteq> 0\<close>, of a, symmetric]
finally show ?thesis .
qed simp
@@ -721,11 +721,11 @@
lemma minus_monom: "- monom a n = monom (-a) n"
by (rule poly_eqI) simp
-lemma coeff_setsum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)"
+lemma coeff_sum: "coeff (\<Sum>x\<in>A. p x) i = (\<Sum>x\<in>A. coeff (p x) i)"
by (cases "finite A", induct set: finite, simp_all)
-lemma monom_setsum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)"
- by (rule poly_eqI) (simp add: coeff_setsum)
+lemma monom_sum: "monom (\<Sum>x\<in>A. a x) n = (\<Sum>x\<in>A. monom (a x) n)"
+ by (rule poly_eqI) (simp add: coeff_sum)
fun plus_coeffs :: "'a::comm_monoid_add list \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
@@ -784,15 +784,15 @@
shows "poly (p - q) x = poly p x - poly q x"
using poly_add [of p "- q" x] by simp
-lemma poly_setsum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
+lemma poly_sum: "poly (\<Sum>k\<in>A. p k) x = (\<Sum>k\<in>A. poly (p k) x)"
by (induct A rule: infinite_finite_induct) simp_all
-lemma degree_setsum_le: "finite S \<Longrightarrow> (\<And> p . p \<in> S \<Longrightarrow> degree (f p) \<le> n)
- \<Longrightarrow> degree (setsum f S) \<le> n"
+lemma degree_sum_le: "finite S \<Longrightarrow> (\<And> p . p \<in> S \<Longrightarrow> degree (f p) \<le> n)
+ \<Longrightarrow> degree (sum f S) \<le> n"
proof (induct S rule: finite_induct)
case (insert p S)
- hence "degree (setsum f S) \<le> n" "degree (f p) \<le> n" by auto
- thus ?case unfolding setsum.insert[OF insert(1-2)] by (metis degree_add_le)
+ hence "degree (sum f S) \<le> n" "degree (f p) \<le> n" by auto
+ thus ?case unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le)
qed simp
lemma poly_as_sum_of_monoms':
@@ -802,7 +802,7 @@
have eq: "\<And>i. {..n} \<inter> {i} = (if i \<le> n then {i} else {})"
by auto
show ?thesis
- using n by (simp add: poly_eq_iff coeff_setsum coeff_eq_0 setsum.If_cases eq
+ using n by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq
if_distrib[where f="\<lambda>x. x * a" for a])
qed
@@ -973,8 +973,8 @@
case 0 show ?case by simp
next
case (pCons a p n) thus ?case
- by (cases n, simp, simp add: setsum_atMost_Suc_shift
- del: setsum_atMost_Suc)
+ by (cases n, simp, simp add: sum_atMost_Suc_shift
+ del: sum_atMost_Suc)
qed
lemma degree_mult_le: "degree (p * q) \<le> degree p + degree q"
@@ -1057,10 +1057,10 @@
lemma poly_setprod: "poly (\<Prod>k\<in>A. p k) x = (\<Prod>k\<in>A. poly (p k) x)"
by (induct A rule: infinite_finite_induct) simp_all
-lemma degree_setprod_setsum_le: "finite S \<Longrightarrow> degree (setprod f S) \<le> setsum (degree o f) S"
+lemma degree_setprod_sum_le: "finite S \<Longrightarrow> degree (setprod f S) \<le> sum (degree o f) S"
proof (induct S rule: finite_induct)
case (insert a S)
- show ?case unfolding setprod.insert[OF insert(1-2)] setsum.insert[OF insert(1-2)]
+ show ?case unfolding setprod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)]
by (rule le_trans[OF degree_mult_le], insert insert, auto)
qed simp
@@ -2852,7 +2852,7 @@
shows "pcompose p [: 0, 1 :] = p"
by (induct p; simp add: pcompose_pCons)
-lemma pcompose_setsum: "pcompose (setsum f A) p = setsum (\<lambda>i. pcompose (f i) p) A"
+lemma pcompose_sum: "pcompose (sum f A) p = sum (\<lambda>i. pcompose (f i) p) A"
by (cases "finite A", induction rule: finite_induct)
(simp_all add: pcompose_1 pcompose_add)
@@ -3185,17 +3185,17 @@
also have "\<dots> = (\<Sum>j\<le>degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))"
unfolding coeff_mult by simp
also have "\<dots> = (\<Sum>j\<in>B. coeff p j * coeff q (degree (p * q) - i - j))"
- by (intro setsum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0)
+ by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0)
also from True have "\<dots> = (\<Sum>j\<in>A. coeff p (degree p - j) * coeff q (degree q - (i - j)))"
- by (intro setsum.reindex_bij_witness[of _ ?f ?f])
+ by (intro sum.reindex_bij_witness[of _ ?f ?f])
(auto simp: A_def B_def degree_mult_eq add_ac)
also have "\<dots> = (\<Sum>j\<le>i. if j \<in> {i - degree q..degree p} then
coeff p (degree p - j) * coeff q (degree q - (i - j)) else 0)"
- by (subst setsum.inter_restrict [symmetric]) (simp_all add: A_def)
+ by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def)
also have "\<dots> = coeff (reflect_poly p * reflect_poly q) i"
- by (fastforce simp: coeff_mult coeff_reflect_poly intro!: setsum.cong)
+ by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong)
finally show ?thesis .
- qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: setsum.neutral)
+ qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral)
qed
qed auto
@@ -3400,7 +3400,7 @@
proof (induct as rule: infinite_finite_induct)
case (insert a as)
hence id: "setprod f (insert a as) = f a * setprod f as"
- "\<And> g. setsum g (insert a as) = g a + setsum g as"
+ "\<And> g. sum g (insert a as) = g a + sum g as"
"insert a as - {a} = as"
by auto
{
@@ -3412,8 +3412,8 @@
by (subst setprod.insert, insert insert, auto)
} note id2 = this
show ?case
- unfolding id pderiv_mult insert(3) setsum_distrib_left
- by (auto simp add: ac_simps id2 intro!: setsum.cong)
+ unfolding id pderiv_mult insert(3) sum_distrib_left
+ by (auto simp add: ac_simps id2 intro!: sum.cong)
qed auto
lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"
@@ -3877,8 +3877,8 @@
have "coeff (monom c n * p) k = (\<Sum>i\<le>k. (if n = i then c else 0) * coeff p (k - i))"
by (simp add: coeff_mult)
also have "\<dots> = (\<Sum>i\<le>k. (if n = i then c * coeff p (k - i) else 0))"
- by (intro setsum.cong) simp_all
- also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))" by (simp add: setsum.delta')
+ by (intro sum.cong) simp_all
+ also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))" by (simp add: sum.delta')
finally show ?thesis .
qed
--- a/src/HOL/Library/Polynomial_FPS.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Polynomial_FPS.thy Mon Oct 17 14:37:32 2016 +0200
@@ -61,7 +61,7 @@
"fps_of_poly (smult c p) = fps_const c * fps_of_poly p"
using fps_of_poly_mult[of "[:c:]" p] by (simp add: fps_of_poly_mult fps_of_poly_const)
-lemma fps_of_poly_setsum: "fps_of_poly (setsum f A) = setsum (\<lambda>x. fps_of_poly (f x)) A"
+lemma fps_of_poly_sum: "fps_of_poly (sum f A) = sum (\<lambda>x. fps_of_poly (f x)) A"
by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_add)
lemma fps_of_poly_sum_list: "fps_of_poly (sum_list xs) = sum_list (map fps_of_poly xs)"
@@ -144,7 +144,7 @@
lemmas fps_of_poly_simps =
fps_of_poly_0 fps_of_poly_1 fps_of_poly_numeral fps_of_poly_const fps_of_poly_X
fps_of_poly_add fps_of_poly_diff fps_of_poly_uminus fps_of_poly_mult fps_of_poly_smult
- fps_of_poly_setsum fps_of_poly_sum_list fps_of_poly_setprod fps_of_poly_prod_list
+ fps_of_poly_sum fps_of_poly_sum_list fps_of_poly_setprod fps_of_poly_prod_list
fps_of_poly_pCons fps_of_poly_pderiv fps_of_poly_power fps_of_poly_monom
fps_of_poly_divide_numeral
--- a/src/HOL/Library/Polynomial_Factorial.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Polynomial_Factorial.thy Mon Oct 17 14:37:32 2016 +0200
@@ -663,7 +663,7 @@
"fract_poly (p - q) = fract_poly p - fract_poly q"
by (intro poly_eqI) (simp_all add: coeff_map_poly)
-lemma to_fract_setsum [simp]: "to_fract (setsum f A) = setsum (\<lambda>x. to_fract (f x)) A"
+lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\<lambda>x. to_fract (f x)) A"
by (cases "finite A", induction A rule: finite_induct) simp_all
lemma fract_poly_mult [simp]:
@@ -851,10 +851,10 @@
also have "(\<Sum>k\<in>\<dots>. coeff a k * coeff b (i + m - k)) =
coeff a i * coeff b m + (\<Sum>k\<in>?A. coeff a k * coeff b (i + m - k))"
(is "_ = _ + ?S")
- by (subst setsum.insert) simp_all
+ by (subst sum.insert) simp_all
finally have eq: "c dvd coeff a i * coeff b m + ?S" .
moreover have "c dvd ?S"
- proof (rule dvd_setsum)
+ proof (rule dvd_sum)
fix k assume k: "k \<in> {..i+m} - {i}"
show "c dvd coeff a k * coeff b (i + m - k)"
proof (cases "k < i")
--- a/src/HOL/Library/Product_Plus.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Product_Plus.thy Mon Oct 17 14:37:32 2016 +0200
@@ -106,7 +106,7 @@
instance prod :: (ab_group_add, ab_group_add) ab_group_add
by standard (simp_all add: prod_eq_iff)
-lemma fst_setsum: "fst (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. fst (f x))"
+lemma fst_sum: "fst (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. fst (f x))"
proof (cases "finite A")
case True
then show ?thesis by induct simp_all
@@ -115,7 +115,7 @@
then show ?thesis by simp
qed
-lemma snd_setsum: "snd (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. snd (f x))"
+lemma snd_sum: "snd (\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. snd (f x))"
proof (cases "finite A")
case True
then show ?thesis by induct simp_all
@@ -124,7 +124,7 @@
then show ?thesis by simp
qed
-lemma setsum_prod: "(\<Sum>x\<in>A. (f x, g x)) = (\<Sum>x\<in>A. f x, \<Sum>x\<in>A. g x)"
+lemma sum_prod: "(\<Sum>x\<in>A. (f x, g x)) = (\<Sum>x\<in>A. f x, \<Sum>x\<in>A. g x)"
proof (cases "finite A")
case True
then show ?thesis by induct (simp_all add: zero_prod_def)
--- a/src/HOL/Library/RBT_Set.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/RBT_Set.thy Mon Oct 17 14:37:32 2016 +0200
@@ -85,7 +85,7 @@
"Pow = Pow" ..
lemma [code, code del]:
- "setsum = setsum" ..
+ "sum = sum" ..
lemma [code, code del]:
"setprod = setprod" ..
@@ -673,13 +673,13 @@
"card (Set t) = fold_keys (\<lambda>_ n. n + 1) t 0"
by (auto simp add: card.eq_fold intro: finite_fold_fold_keys comp_fun_commute_const)
-lemma setsum_Set [code]:
- "setsum f (Set xs) = fold_keys (plus o f) xs 0"
+lemma sum_Set [code]:
+ "sum f (Set xs) = fold_keys (plus o f) xs 0"
proof -
have "comp_fun_commute (\<lambda>x. op + (f x))"
by standard (auto simp: ac_simps)
then show ?thesis
- by (auto simp add: setsum.eq_fold finite_fold_fold_keys o_def)
+ by (auto simp add: sum.eq_fold finite_fold_fold_keys o_def)
qed
lemma the_elem_set [code]:
--- a/src/HOL/Library/Set_Algebras.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Set_Algebras.thy Mon Oct 17 14:37:32 2016 +0200
@@ -347,24 +347,24 @@
lemma finite_set_times: "finite s \<Longrightarrow> finite t \<Longrightarrow> finite (s * t)"
by (simp add: set_times_image)
-lemma set_setsum_alt:
+lemma set_sum_alt:
assumes fin: "finite I"
- shows "setsum S I = {setsum s I |s. \<forall>i\<in>I. s i \<in> S i}"
- (is "_ = ?setsum I")
+ shows "sum S I = {sum s I |s. \<forall>i\<in>I. s i \<in> S i}"
+ (is "_ = ?sum I")
using fin
proof induct
case empty
then show ?case by simp
next
case (insert x F)
- have "setsum S (insert x F) = S x + ?setsum F"
+ have "sum S (insert x F) = S x + ?sum F"
using insert.hyps by auto
- also have "\<dots> = {s x + setsum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}"
+ also have "\<dots> = {s x + sum s F |s. \<forall> i\<in>insert x F. s i \<in> S i}"
unfolding set_plus_def
proof safe
fix y s
assume "y \<in> S x" "\<forall>i\<in>F. s i \<in> S i"
- then show "\<exists>s'. y + setsum s F = s' x + setsum s' F \<and> (\<forall>i\<in>insert x F. s' i \<in> S i)"
+ then show "\<exists>s'. y + sum s F = s' x + sum s' F \<and> (\<forall>i\<in>insert x F. s' i \<in> S i)"
using insert.hyps
by (intro exI[of _ "\<lambda>i. if i \<in> F then s i else y"]) (auto simp add: set_plus_def)
qed auto
@@ -372,12 +372,12 @@
using insert.hyps by auto
qed
-lemma setsum_set_cond_linear:
+lemma sum_set_cond_linear:
fixes f :: "'a::comm_monoid_add set \<Rightarrow> 'b::comm_monoid_add set"
assumes [intro!]: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> P (A + B)" "P {0}"
and f: "\<And>A B. P A \<Longrightarrow> P B \<Longrightarrow> f (A + B) = f A + f B" "f {0} = {0}"
assumes all: "\<And>i. i \<in> I \<Longrightarrow> P (S i)"
- shows "f (setsum S I) = setsum (f \<circ> S) I"
+ shows "f (sum S I) = sum (f \<circ> S) I"
proof (cases "finite I")
case True
from this all show ?thesis
@@ -386,7 +386,7 @@
then show ?case by (auto intro!: f)
next
case (insert x F)
- from \<open>finite F\<close> \<open>\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)\<close> have "P (setsum S F)"
+ from \<open>finite F\<close> \<open>\<And>i. i \<in> insert x F \<Longrightarrow> P (S i)\<close> have "P (sum S F)"
by induct auto
with insert show ?case
by (simp, subst f) auto
@@ -396,11 +396,11 @@
then show ?thesis by (auto intro!: f)
qed
-lemma setsum_set_linear:
+lemma sum_set_linear:
fixes f :: "'a::comm_monoid_add set \<Rightarrow> 'b::comm_monoid_add set"
assumes "\<And>A B. f(A) + f(B) = f(A + B)" "f {0} = {0}"
- shows "f (setsum S I) = setsum (f \<circ> S) I"
- using setsum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto
+ shows "f (sum S I) = sum (f \<circ> S) I"
+ using sum_set_cond_linear[of "\<lambda>x. True" f I S] assms by auto
lemma set_times_Un_distrib:
"A * (B \<union> C) = A * B \<union> A * C"
--- a/src/HOL/Library/Stirling.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Library/Stirling.thy Mon Oct 17 14:37:32 2016 +0200
@@ -110,7 +110,7 @@
also have "\<dots> = Suc n * (\<Sum>k=1..n. fact n div k) + Suc n * fact n div Suc n"
by (metis nat.distinct(1) nonzero_mult_div_cancel_left)
also have "\<dots> = (\<Sum>k=1..n. fact (Suc n) div k) + fact (Suc n) div Suc n"
- by (simp add: setsum_distrib_left div_mult_swap dvd_fact)
+ by (simp add: sum_distrib_left div_mult_swap dvd_fact)
also have "\<dots> = (\<Sum>k=1..Suc n. fact (Suc n) div k)"
by simp
finally show ?thesis .
@@ -149,24 +149,24 @@
qed
qed
-lemma setsum_stirling: "(\<Sum>k\<le>n. stirling n k) = fact n"
+lemma sum_stirling: "(\<Sum>k\<le>n. stirling n k) = fact n"
proof (induct n)
case 0
then show ?case by simp
next
case (Suc n)
have "(\<Sum>k\<le>Suc n. stirling (Suc n) k) = stirling (Suc n) 0 + (\<Sum>k\<le>n. stirling (Suc n) (Suc k))"
- by (simp only: setsum_atMost_Suc_shift)
+ by (simp only: sum_atMost_Suc_shift)
also have "\<dots> = (\<Sum>k\<le>n. stirling (Suc n) (Suc k))"
by simp
also have "\<dots> = (\<Sum>k\<le>n. n * stirling n (Suc k) + stirling n k)"
by simp
also have "\<dots> = n * (\<Sum>k\<le>n. stirling n (Suc k)) + (\<Sum>k\<le>n. stirling n k)"
- by (simp add: setsum.distrib setsum_distrib_left)
+ by (simp add: sum.distrib sum_distrib_left)
also have "\<dots> = n * fact n + fact n"
proof -
have "n * (\<Sum>k\<le>n. stirling n (Suc k)) = n * ((\<Sum>k\<le>Suc n. stirling n k) - stirling n 0)"
- by (metis add_diff_cancel_left' setsum_atMost_Suc_shift)
+ by (metis add_diff_cancel_left' sum_atMost_Suc_shift)
also have "\<dots> = n * (\<Sum>k\<le>n. stirling n k)"
by (cases n) simp_all
also have "\<dots> = n * fact n"
@@ -192,10 +192,10 @@
(of_nat (n * stirling n 0) * x ^ 0 +
(\<Sum>i\<le>n. of_nat (n * stirling n (Suc i)) * (x ^ Suc i))) +
(\<Sum>i\<le>n. of_nat (stirling n i) * (x ^ Suc i))"
- by (subst setsum_atMost_Suc_shift) (simp add: setsum.distrib ring_distribs)
+ by (subst sum_atMost_Suc_shift) (simp add: sum.distrib ring_distribs)
also have "\<dots> = pochhammer x (Suc n)"
- by (subst setsum_atMost_Suc_shift [symmetric])
- (simp add: algebra_simps setsum.distrib setsum_distrib_left pochhammer_Suc Suc [symmetric])
+ by (subst sum_atMost_Suc_shift [symmetric])
+ (simp add: algebra_simps sum.distrib sum_distrib_left pochhammer_Suc Suc [symmetric])
finally show ?case .
qed
--- a/src/HOL/Lifting_Set.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Lifting_Set.thy Mon Oct 17 14:37:32 2016 +0200
@@ -285,7 +285,7 @@
by (simp add: reindex_bij_betw)
qed
-lemmas setsum_parametric = setsum.F_parametric
+lemmas sum_parametric = sum.F_parametric
lemmas setprod_parametric = setprod.F_parametric
lemma rel_set_UNION:
--- a/src/HOL/Limits.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Limits.thy Mon Oct 17 14:37:32 2016 +0200
@@ -641,20 +641,20 @@
shows "(f \<longlongrightarrow> 0) F \<Longrightarrow> (g \<longlongrightarrow> 0) F \<Longrightarrow> ((\<lambda>x. f x + g x) \<longlongrightarrow> 0) F"
by (drule (1) tendsto_add) simp
-lemma tendsto_setsum [tendsto_intros]:
+lemma tendsto_sum [tendsto_intros]:
fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> 'c::topological_comm_monoid_add"
shows "(\<And>i. i \<in> I \<Longrightarrow> (f i \<longlongrightarrow> a i) F) \<Longrightarrow> ((\<lambda>x. \<Sum>i\<in>I. f i x) \<longlongrightarrow> (\<Sum>i\<in>I. a i)) F"
by (induct I rule: infinite_finite_induct) (simp_all add: tendsto_add)
-lemma continuous_setsum [continuous_intros]:
+lemma continuous_sum [continuous_intros]:
fixes f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add"
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous F (f i)) \<Longrightarrow> continuous F (\<lambda>x. \<Sum>i\<in>I. f i x)"
- unfolding continuous_def by (rule tendsto_setsum)
-
-lemma continuous_on_setsum [continuous_intros]:
+ unfolding continuous_def by (rule tendsto_sum)
+
+lemma continuous_on_sum [continuous_intros]:
fixes f :: "'a \<Rightarrow> 'b::topological_space \<Rightarrow> 'c::topological_comm_monoid_add"
shows "(\<And>i. i \<in> I \<Longrightarrow> continuous_on S (f i)) \<Longrightarrow> continuous_on S (\<lambda>x. \<Sum>i\<in>I. f i x)"
- unfolding continuous_on_def by (auto intro: tendsto_setsum)
+ unfolding continuous_on_def by (auto intro: tendsto_sum)
instance nat :: topological_comm_monoid_add
by standard
@@ -1886,7 +1886,7 @@
shows "convergent (\<lambda>n. X n + Y n)"
using assms unfolding convergent_def by (blast intro: tendsto_add)
-lemma convergent_setsum:
+lemma convergent_sum:
fixes X :: "'a \<Rightarrow> nat \<Rightarrow> 'b::real_normed_vector"
shows "(\<And>i. i \<in> A \<Longrightarrow> convergent (\<lambda>n. X i n)) \<Longrightarrow> convergent (\<lambda>n. \<Sum>i\<in>A. X i n)"
by (induct A rule: infinite_finite_induct) (simp_all add: convergent_const convergent_add)
@@ -2299,9 +2299,9 @@
for f :: "'a::t2_space \<Rightarrow> 'b::{power,real_normed_algebra}"
by (fact continuous_power)
-lemma isCont_setsum [simp]: "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
+lemma isCont_sum [simp]: "\<forall>i\<in>A. isCont (f i) a \<Longrightarrow> isCont (\<lambda>x. \<Sum>i\<in>A. f i x) a"
for f :: "'a \<Rightarrow> 'b::t2_space \<Rightarrow> 'c::topological_comm_monoid_add"
- by (auto intro: continuous_setsum)
+ by (auto intro: continuous_sum)
subsection \<open>Uniform Continuity\<close>
--- a/src/HOL/List.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/List.thy Mon Oct 17 14:37:32 2016 +0200
@@ -3808,12 +3808,12 @@
lemma count_le_length: "count_list xs x \<le> length xs"
by (induction xs) auto
-lemma setsum_count_set:
- "set xs \<subseteq> X \<Longrightarrow> finite X \<Longrightarrow> setsum (count_list xs) X = length xs"
+lemma sum_count_set:
+ "set xs \<subseteq> X \<Longrightarrow> finite X \<Longrightarrow> sum (count_list xs) X = length xs"
apply(induction xs arbitrary: X)
apply simp
-apply (simp add: setsum.If_cases)
-by (metis Diff_eq setsum.remove)
+apply (simp add: sum.If_cases)
+by (metis Diff_eq sum.remove)
subsubsection \<open>@{const List.extract}\<close>
--- a/src/HOL/MacLaurin.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/MacLaurin.thy Mon Oct 17 14:37:32 2016 +0200
@@ -48,7 +48,7 @@
unfolding atLeast0LessThan[symmetric] by auto
have "(\<Sum>x<n - m. real x * t ^ (x - Suc 0) * diff (m + x) 0 / fact x) =
(\<Sum>x<n - Suc m. real (Suc x) * t ^ x * diff (Suc m + x) 0 / fact (Suc x))"
- unfolding intvl by (subst setsum.insert) (auto simp add: setsum.reindex)
+ unfolding intvl by (subst sum.insert) (auto simp add: sum.reindex)
moreover
have fact_neq_0: "\<And>x. (fact x) + real x * (fact x) \<noteq> 0"
by (metis add_pos_pos fact_gt_zero less_add_same_cancel1 less_add_same_cancel2
@@ -69,7 +69,7 @@
and diff_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
shows
"\<exists>t::real. 0 < t \<and> t < h \<and>
- f h = setsum (\<lambda>m. (diff m 0 / fact m) * h ^ m) {..<n} + (diff n t / fact n) * h ^ n"
+ f h = sum (\<lambda>m. (diff m 0 / fact m) * h ^ m) {..<n} + (diff n t / fact n) * h ^ n"
proof -
from n obtain m where m: "n = Suc m"
by (cases n) (simp add: n)
@@ -79,19 +79,19 @@
using Maclaurin_lemma [OF h] ..
define g where [abs_def]: "g t =
- f t - (setsum (\<lambda>m. (diff m 0 / fact m) * t^m) {..<n} + B * (t^n / fact n))" for t
+ f t - (sum (\<lambda>m. (diff m 0 / fact m) * t^m) {..<n} + B * (t^n / fact n))" for t
have g2: "g 0 = 0" "g h = 0"
- by (simp_all add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
+ by (simp_all add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 sum.reindex)
define difg where [abs_def]: "difg m t =
- diff m t - (setsum (\<lambda>p. (diff (m + p) 0 / fact p) * (t ^ p)) {..<n-m} +
+ diff m t - (sum (\<lambda>p. (diff (m + p) 0 / fact p) * (t ^ p)) {..<n-m} +
B * ((t ^ (n - m)) / fact (n - m)))" for m t
have difg_0: "difg 0 = g"
by (simp add: difg_def g_def diff_0)
have difg_Suc: "\<forall>m t. m < n \<and> 0 \<le> t \<and> t \<le> h \<longrightarrow> DERIV (difg m) t :> difg (Suc m) t"
using diff_Suc m unfolding difg_def [abs_def] by (rule Maclaurin_lemma2)
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)
+ by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff sum.reindex)
have isCont_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> isCont (difg m) x"
by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
have differentiable_difg: "\<And>m x. m < n \<Longrightarrow> 0 \<le> x \<Longrightarrow> x \<le> h \<Longrightarrow> difg m differentiable (at x)"
@@ -198,7 +198,7 @@
by (auto simp: power_mult_distrib[symmetric])
moreover
have "(\<Sum>m<n. (- 1) ^ m * diff m 0 * (- h) ^ m / fact m) = (\<Sum>m<n. diff m 0 * h ^ m / fact m)"
- by (auto intro: setsum.cong simp add: power_mult_distrib[symmetric])
+ by (auto intro: sum.cong simp add: power_mult_distrib[symmetric])
ultimately have "h < - t \<and> - t < 0 \<and>
f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n (- t) / (fact n) * h ^ n"
by auto
@@ -394,7 +394,7 @@
apply simp
apply (simp add: sin_expansion_lemma del: of_nat_Suc)
apply (force intro!: derivative_eq_intros)
- apply (subst (asm) setsum.neutral; auto)
+ apply (subst (asm) sum.neutral; auto)
apply (rule ccontr)
apply simp
apply (drule_tac x = x in spec)
@@ -402,7 +402,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
done
@@ -423,7 +423,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
done
@@ -440,7 +440,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp add: sin_coeff_def sin_zero_iff elim: oddE simp del: of_nat_Suc)
done
@@ -463,7 +463,7 @@
apply (simp add: cos_expansion_lemma del: of_nat_Suc)
apply (cases n)
apply simp
- apply (simp del: setsum_lessThan_Suc)
+ apply (simp del: sum_lessThan_Suc)
apply (rule ccontr)
apply simp
apply (drule_tac x = x in spec)
@@ -471,7 +471,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp add: cos_coeff_def cos_zero_iff elim: evenE)
done
@@ -487,7 +487,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
done
@@ -503,7 +503,7 @@
apply (erule ssubst)
apply (rule_tac x = t in exI)
apply simp
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (auto simp: cos_coeff_def cos_zero_iff elim: evenE)
done
@@ -548,7 +548,7 @@
unfolding sin_coeff_def
apply (subst t2)
apply (rule sin_bound_lemma)
- apply (rule setsum.cong[OF refl])
+ apply (rule sum.cong[OF refl])
apply (subst diff_m_0, simp)
using est
apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
--- a/src/HOL/Metis_Examples/Big_O.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Metis_Examples/Big_O.thy Mon Oct 17 14:37:32 2016 +0200
@@ -489,78 +489,78 @@
apply (simp add: fun_diff_def)
done
-subsection \<open>Setsum\<close>
+subsection \<open>Sum\<close>
-lemma bigo_setsum_main: "\<forall>x. \<forall>y \<in> A x. 0 <= h x y \<Longrightarrow>
+lemma bigo_sum_main: "\<forall>x. \<forall>y \<in> A x. 0 <= h x y \<Longrightarrow>
\<exists>c. \<forall>x. \<forall>y \<in> A x. \<bar>f x y\<bar> <= c * (h x y) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)"
apply (auto simp add: bigo_def)
apply (rule_tac x = "\<bar>c\<bar>" in exI)
apply (subst abs_of_nonneg) back back
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply force
-apply (subst setsum_distrib_left)
+apply (subst sum_distrib_left)
apply (rule allI)
apply (rule order_trans)
- apply (rule setsum_abs)
-apply (rule setsum_mono)
+ apply (rule sum_abs)
+apply (rule sum_mono)
by (metis abs_ge_self abs_mult_pos order_trans)
-lemma bigo_setsum1: "\<forall>x y. 0 <= h x y \<Longrightarrow>
+lemma bigo_sum1: "\<forall>x y. 0 <= h x y \<Longrightarrow>
\<exists>c. \<forall>x y. \<bar>f x y\<bar> <= c * (h x y) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f x y) =o O(\<lambda>x. \<Sum>y \<in> A x. h x y)"
-by (metis (no_types) bigo_setsum_main)
+by (metis (no_types) bigo_sum_main)
-lemma bigo_setsum2: "\<forall>y. 0 <= h y \<Longrightarrow>
+lemma bigo_sum2: "\<forall>y. 0 <= h y \<Longrightarrow>
\<exists>c. \<forall>y. \<bar>f y\<bar> <= c * (h y) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. f y) =o O(\<lambda>x. \<Sum>y \<in> A x. h y)"
-apply (rule bigo_setsum1)
+apply (rule bigo_sum1)
by metis+
-lemma bigo_setsum3: "f =o O(h) \<Longrightarrow>
+lemma bigo_sum3: "f =o O(h) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h(k x y)\<bar>)"
-apply (rule bigo_setsum1)
+apply (rule bigo_sum1)
apply (rule allI)+
apply (rule abs_ge_zero)
apply (unfold bigo_def)
apply (auto simp add: abs_mult)
by (metis abs_ge_zero mult.left_commute mult_left_mono)
-lemma bigo_setsum4: "f =o g +o O(h) \<Longrightarrow>
+lemma bigo_sum4: "f =o g +o O(h) \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. l x y * f(k x y)) =o
(\<lambda>x. \<Sum>y \<in> A x. l x y * g(k x y)) +o
O(\<lambda>x. \<Sum>y \<in> A x. \<bar>l x y * h(k x y)\<bar>)"
apply (rule set_minus_imp_plus)
apply (subst fun_diff_def)
-apply (subst setsum_subtractf [symmetric])
+apply (subst sum_subtractf [symmetric])
apply (subst right_diff_distrib [symmetric])
-apply (rule bigo_setsum3)
+apply (rule bigo_sum3)
by (metis (lifting, no_types) fun_diff_def set_plus_imp_minus ext)
-lemma bigo_setsum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow>
+lemma bigo_sum5: "f =o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow>
\<forall>x. 0 <= h x \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o
O(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y))"
apply (subgoal_tac "(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y)) =
(\<lambda>x. \<Sum>y \<in> A x. \<bar>(l x y) * h(k x y)\<bar>)")
apply (erule ssubst)
- apply (erule bigo_setsum3)
+ apply (erule bigo_sum3)
apply (rule ext)
-apply (rule setsum.cong)
+apply (rule sum.cong)
apply (rule refl)
by (metis abs_of_nonneg zero_le_mult_iff)
-lemma bigo_setsum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow>
+lemma bigo_sum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow>
\<forall>x. 0 <= h x \<Longrightarrow>
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * f(k x y)) =o
(\<lambda>x. \<Sum>y \<in> A x. (l x y) * g(k x y)) +o
O(\<lambda>x. \<Sum>y \<in> A x. (l x y) * h(k x y))"
apply (rule set_minus_imp_plus)
apply (subst fun_diff_def)
- apply (subst setsum_subtractf [symmetric])
+ apply (subst sum_subtractf [symmetric])
apply (subst right_diff_distrib [symmetric])
- apply (rule bigo_setsum5)
+ apply (rule bigo_sum5)
apply (subst fun_diff_def [symmetric])
apply (drule set_plus_imp_minus)
apply auto
--- a/src/HOL/Nat_Transfer.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Nat_Transfer.thy Mon Oct 17 14:37:32 2016 +0200
@@ -186,13 +186,13 @@
]
-text \<open>setsum and setprod\<close>
+text \<open>sum and setprod\<close>
(* this handles the case where the *domain* of f is nat *)
lemma transfer_nat_int_sum_prod:
- "setsum f A = setsum (%x. f (nat x)) (int ` A)"
+ "sum f A = sum (%x. f (nat x)) (int ` A)"
"setprod f A = setprod (%x. f (nat x)) (int ` A)"
- apply (subst setsum.reindex)
+ apply (subst sum.reindex)
apply (unfold inj_on_def, auto)
apply (subst setprod.reindex)
apply (unfold inj_on_def o_def, auto)
@@ -200,17 +200,17 @@
(* this handles the case where the *range* of f is nat *)
lemma transfer_nat_int_sum_prod2:
- "setsum f A = nat(setsum (%x. int (f x)) A)"
+ "sum f A = nat(sum (%x. int (f x)) A)"
"setprod f A = nat(setprod (%x. int (f x)) A)"
- apply (simp only: int_setsum [symmetric] nat_int)
+ apply (simp only: int_sum [symmetric] nat_int)
apply (simp only: int_setprod [symmetric] nat_int)
done
lemma transfer_nat_int_sum_prod_closure:
- "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
+ "nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> sum f A >= 0"
"nat_set A \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
unfolding nat_set_def
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply auto
apply (rule setprod_nonneg)
apply auto
@@ -222,10 +222,10 @@
also: what does =simp=> do?
lemma transfer_nat_int_sum_prod_closure:
- "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setsum f A >= 0"
+ "(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> sum f A >= 0"
"(!!x. x : A ==> f x >= (0::int)) \<Longrightarrow> setprod f A >= 0"
unfolding nat_set_def simp_implies_def
- apply (rule setsum_nonneg)
+ apply (rule sum_nonneg)
apply auto
apply (rule setprod_nonneg)
apply auto
@@ -233,16 +233,16 @@
*)
(* Making A = B in this lemma doesn't work. Why not?
- Also, why aren't setsum.cong and setprod.cong enough,
+ Also, why aren't sum.cong and setprod.cong enough,
with the previously mentioned rule turned on? *)
lemma transfer_nat_int_sum_prod_cong:
"A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
- setsum f A = setsum g B"
+ sum f A = sum g B"
"A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
setprod f A = setprod g B"
unfolding nat_set_def
- apply (subst setsum.cong, assumption)
+ apply (subst sum.cong, assumption)
apply auto [2]
apply (subst setprod.cong, assumption, auto)
done
@@ -389,13 +389,13 @@
]
-text \<open>setsum and setprod\<close>
+text \<open>sum and setprod\<close>
(* this handles the case where the *domain* of f is int *)
lemma transfer_int_nat_sum_prod:
- "nat_set A \<Longrightarrow> setsum f A = setsum (%x. f (int x)) (nat ` A)"
+ "nat_set A \<Longrightarrow> sum f A = sum (%x. f (int x)) (nat ` A)"
"nat_set A \<Longrightarrow> setprod f A = setprod (%x. f (int x)) (nat ` A)"
- apply (subst setsum.reindex)
+ apply (subst sum.reindex)
apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff)
apply (subst setprod.reindex)
apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff
@@ -404,14 +404,14 @@
(* this handles the case where the *range* of f is int *)
lemma transfer_int_nat_sum_prod2:
- "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> setsum f A = int(setsum (%x. nat (f x)) A)"
+ "(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow> sum f A = int(sum (%x. nat (f x)) A)"
"(!!x. x:A \<Longrightarrow> is_nat (f x)) \<Longrightarrow>
setprod f A = int(setprod (%x. nat (f x)) A)"
unfolding is_nat_def
- by (subst int_setsum) auto
+ by (subst int_sum) auto
declare transfer_morphism_int_nat [transfer add
return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2
- cong: setsum.cong setprod.cong]
+ cong: sum.cong setprod.cong]
end
--- a/src/HOL/Nitpick.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Nitpick.thy Mon Oct 17 14:37:32 2016 +0200
@@ -72,8 +72,8 @@
definition card' :: "'a set \<Rightarrow> nat" where
"card' A \<equiv> if finite A then length (SOME xs. set xs = A \<and> distinct xs) else 0"
-definition setsum' :: "('a \<Rightarrow> 'b::comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
- "setsum' f A \<equiv> if finite A then sum_list (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
+definition sum' :: "('a \<Rightarrow> 'b::comm_monoid_add) \<Rightarrow> 'a set \<Rightarrow> 'b" where
+ "sum' f A \<equiv> if finite A then sum_list (map f (SOME xs. set xs = A \<and> distinct xs)) else 0"
inductive fold_graph' :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool" where
"fold_graph' f z {} z" |
@@ -224,7 +224,7 @@
setup \<open>
Nitpick_HOL.register_ersatz_global
[(@{const_name card}, @{const_name card'}),
- (@{const_name setsum}, @{const_name setsum'}),
+ (@{const_name sum}, @{const_name sum'}),
(@{const_name fold_graph}, @{const_name fold_graph'}),
(@{const_name wf}, @{const_name wf'}),
(@{const_name wf_wfrec}, @{const_name wf_wfrec'}),
@@ -232,14 +232,14 @@
\<close>
hide_const (open) unknown is_unknown bisim bisim_iterator_max Quot safe_The FunBox PairBox Word prod
- refl' wf' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac
+ refl' wf' card' sum' fold_graph' nat_gcd nat_lcm int_gcd int_lcm Frac Abs_Frac Rep_Frac
zero_frac one_frac num denom norm_frac frac plus_frac times_frac uminus_frac number_of_frac
inverse_frac less_frac less_eq_frac of_frac wf_wfrec wf_wfrec wfrec'
hide_type (open) bisim_iterator fun_box pair_box unsigned_bit signed_bit word
hide_fact (open) Ex1_unfold rtrancl_unfold rtranclp_unfold tranclp_unfold prod_def refl'_def wf'_def
- card'_def setsum'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
+ card'_def sum'_def The_psimp Eps_psimp case_unit_unfold case_nat_unfold
size_list_simp nat_lcm_def int_gcd_def int_lcm_def Frac_def zero_frac_def one_frac_def
num_def denom_def frac_def plus_frac_def times_frac_def uminus_frac_def
number_of_frac_def inverse_frac_def less_frac_def less_eq_frac_def of_frac_def wf_wfrec'_def
--- a/src/HOL/Nonstandard_Analysis/HSeries.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Nonstandard_Analysis/HSeries.thy Mon Oct 17 14:37:32 2016 +0200
@@ -14,11 +14,11 @@
definition
sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal" where
"sumhr =
- (%(M,N,f). starfun2 (%m n. setsum f {m..<n}) M N)"
+ (%(M,N,f). starfun2 (%m n. sum f {m..<n}) M N)"
definition
NSsums :: "[nat=>real,real] => bool" (infixr "NSsums" 80) where
- "f NSsums s = (%n. setsum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
+ "f NSsums s = (%n. sum f {..<n}) \<longlonglongrightarrow>\<^sub>N\<^sub>S s"
definition
NSsummable :: "(nat=>real) => bool" where
@@ -28,7 +28,7 @@
NSsuminf :: "(nat=>real) => real" where
"NSsuminf f = (THE s. f NSsums s)"
-lemma sumhr_app: "sumhr(M,N,f) = ( *f2* (\<lambda>m n. setsum f {m..<n})) M N"
+lemma sumhr_app: "sumhr(M,N,f) = ( *f2* (\<lambda>m n. sum f {m..<n})) M N"
by (simp add: sumhr_def)
text\<open>Base case in definition of @{term sumr}\<close>
@@ -55,21 +55,21 @@
lemma sumhr_add:
"!!m n. sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)"
-unfolding sumhr_app by transfer (rule setsum.distrib [symmetric])
+unfolding sumhr_app by transfer (rule sum.distrib [symmetric])
lemma sumhr_mult:
"!!m n. hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
-unfolding sumhr_app by transfer (rule setsum_distrib_left)
+unfolding sumhr_app by transfer (rule sum_distrib_left)
lemma sumhr_split_add:
"!!n p. n < p ==> sumhr(0,n,f) + sumhr(n,p,f) = sumhr(0,p,f)"
-unfolding sumhr_app by transfer (simp add: setsum_add_nat_ivl)
+unfolding sumhr_app by transfer (simp add: sum_add_nat_ivl)
lemma sumhr_split_diff: "n<p ==> sumhr(0,p,f) - sumhr(0,n,f) = sumhr(n,p,f)"
by (drule_tac f = f in sumhr_split_add [symmetric], simp)
lemma sumhr_hrabs: "!!m n. \<bar>sumhr(m,n,f)\<bar> \<le> sumhr(m,n,%i. \<bar>f i\<bar>)"
-unfolding sumhr_app by transfer (rule setsum_abs)
+unfolding sumhr_app by transfer (rule sum_abs)
text\<open>other general version also needed\<close>
lemma sumhr_fun_hypnat_eq:
@@ -86,12 +86,12 @@
unfolding sumhr_app by transfer simp
lemma sumhr_minus: "!!m n. sumhr(m, n, %i. - f i) = - sumhr(m, n, f)"
-unfolding sumhr_app by transfer (rule setsum_negf)
+unfolding sumhr_app by transfer (rule sum_negf)
lemma sumhr_shift_bounds:
"!!m n. sumhr(m+hypnat_of_nat k,n+hypnat_of_nat k,f) =
sumhr(m,n,%i. f(i + k))"
-unfolding sumhr_app by transfer (rule setsum_shift_bounds_nat_ivl)
+unfolding sumhr_app by transfer (rule sum_shift_bounds_nat_ivl)
subsection\<open>Nonstandard Sums\<close>
@@ -126,7 +126,7 @@
(hypreal_of_nat (na - m) * hypreal_of_real r)"
unfolding sumhr_app by transfer simp
-lemma starfunNat_sumr: "!!N. ( *f* (%n. setsum f {0..<n})) N = sumhr(0,N,f)"
+lemma starfunNat_sumr: "!!N. ( *f* (%n. sum f {0..<n})) N = sumhr(0,N,f)"
unfolding sumhr_app by transfer (rule refl)
lemma sumhr_hrabs_approx [simp]: "sumhr(0, M, f) \<approx> sumhr(0, N, f)
@@ -162,7 +162,7 @@
by (simp add: suminf_NSsuminf_iff [symmetric] sums_NSsums_iff sums_unique)
lemma NSseries_zero:
- "\<forall>m. n \<le> Suc m --> f(m) = 0 ==> f NSsums (setsum f {..<n})"
+ "\<forall>m. n \<le> Suc m --> f(m) = 0 ==> f NSsums (sum f {..<n})"
by (auto simp add: sums_NSsums_iff [symmetric] not_le[symmetric] intro!: sums_finite)
lemma NSsummable_NSCauchy:
--- a/src/HOL/Nonstandard_Analysis/Nonstandard_Analysis.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Nonstandard_Analysis/Nonstandard_Analysis.thy Mon Oct 17 14:37:32 2016 +0200
@@ -10,4 +10,4 @@
imports Hypercomplex
begin
-end
\ No newline at end of file
+end
--- a/src/HOL/NthRoot.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/NthRoot.thy Mon Oct 17 14:37:32 2016 +0200
@@ -720,7 +720,7 @@
by (simp add: x_def)
also have "\<dots> \<le> (\<Sum>k=0..n. of_nat (n choose k) * x n^k)"
using \<open>2 < n\<close>
- by (intro setsum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
+ by (intro sum_mono2) (auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
also have "\<dots> = (x n + 1) ^ n"
by (simp add: binomial_ring)
also have "\<dots> = n"
@@ -761,7 +761,7 @@
by (simp add: x_def)
also have "\<dots> \<le> (\<Sum>k=0..n. of_nat (n choose k) * x n^k)"
using \<open>1 < n\<close> \<open>1 \<le> c\<close>
- by (intro setsum_mono2)
+ by (intro sum_mono2)
(auto intro!: mult_nonneg_nonneg zero_le_power simp: x_def le_diff_eq)
also have "\<dots> = (x n + 1) ^ n"
by (simp add: binomial_ring)
--- a/src/HOL/Number_Theory/Cong.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Number_Theory/Cong.thy Mon Oct 17 14:37:32 2016 +0200
@@ -177,7 +177,7 @@
lemma cong_exp_int: "[(x::int) = y] (mod n) \<Longrightarrow> [x^k = y^k] (mod n)"
by (induct k) (auto simp add: cong_mult_int)
-lemma cong_setsum_nat [rule_format]:
+lemma cong_sum_nat [rule_format]:
"(\<forall>x\<in>A. [((f x)::nat) = g x] (mod m)) \<longrightarrow>
[(\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. g x)] (mod m)"
apply (cases "finite A")
@@ -185,7 +185,7 @@
apply (auto intro: cong_add_nat)
done
-lemma cong_setsum_int [rule_format]:
+lemma cong_sum_int [rule_format]:
"(\<forall>x\<in>A. [((f x)::int) = g x] (mod m)) \<longrightarrow>
[(\<Sum>x\<in>A. f x) = (\<Sum>x\<in>A. g x)] (mod m)"
apply (cases "finite A")
@@ -794,7 +794,7 @@
proof -
from fin a have "?x = (\<Sum>j \<in> {i}. u j * b j) +
(\<Sum>j \<in> A - {i}. u j * b j)"
- by (subst setsum.union_disjoint [symmetric], auto intro: setsum.cong)
+ by (subst sum.union_disjoint [symmetric], auto intro: sum.cong)
then have "[?x = u i * b i + (\<Sum>j \<in> A - {i}. u j * b j)] (mod m i)"
by auto
also have "[u i * b i + (\<Sum>j \<in> A - {i}. u j * b j) =
@@ -802,7 +802,7 @@
apply (rule cong_add_nat)
apply (rule cong_scalar2_nat)
using bprop a apply blast
- apply (rule cong_setsum_nat)
+ apply (rule cong_sum_nat)
apply (rule cong_scalar2_nat)
using bprop apply auto
apply (rule cong_dvd_modulus_nat)
--- a/src/HOL/Number_Theory/Factorial_Ring.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Number_Theory/Factorial_Ring.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1300,7 +1300,7 @@
proof -
have "multiplicity p (setprod f A) = (\<Sum>x\<in>#mset_set A. multiplicity p (f x))"
using assms by (subst setprod_unfold_prod_mset)
- (simp_all add: prime_elem_multiplicity_prod_mset_distrib setsum_unfold_sum_mset
+ (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset
multiset.map_comp o_def)
also from \<open>finite A\<close> have "\<dots> = (\<Sum>x\<in>A. multiplicity p (f x))"
by (induction A rule: finite_induct) simp_all
--- a/src/HOL/Number_Theory/Fib.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Number_Theory/Fib.thy Mon Oct 17 14:37:32 2016 +0200
@@ -100,18 +100,18 @@
\<comment> \<open>Law 6.111\<close>
by (induct m n rule: gcd_nat_induct) (simp_all add: gcd_non_0_nat gcd.commute gcd_fib_mod)
-theorem fib_mult_eq_setsum_nat: "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
+theorem fib_mult_eq_sum_nat: "fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
by (induct n rule: nat.induct) (auto simp add: field_simps)
subsection \<open>Fibonacci and Binomial Coefficients\<close>
-lemma setsum_drop_zero: "(\<Sum>k = 0..Suc n. if 0<k then (f (k - 1)) else 0) = (\<Sum>j = 0..n. f j)"
+lemma sum_drop_zero: "(\<Sum>k = 0..Suc n. if 0<k then (f (k - 1)) else 0) = (\<Sum>j = 0..n. f j)"
by (induct n) auto
-lemma setsum_choose_drop_zero:
+lemma sum_choose_drop_zero:
"(\<Sum>k = 0..Suc n. if k=0 then 0 else (Suc n - k) choose (k - 1)) = (\<Sum>j = 0..n. (n-j) choose j)"
- by (rule trans [OF setsum.cong setsum_drop_zero]) auto
+ by (rule trans [OF sum.cong sum_drop_zero]) auto
lemma ne_diagonal_fib: "(\<Sum>k = 0..n. (n-k) choose k) = fib (Suc n)"
proof (induct n rule: fib.induct)
@@ -124,13 +124,13 @@
case (3 n)
have "(\<Sum>k = 0..Suc n. Suc (Suc n) - k choose k) =
(\<Sum>k = 0..Suc n. (Suc n - k choose k) + (if k=0 then 0 else (Suc n - k choose (k - 1))))"
- by (rule setsum.cong) (simp_all add: choose_reduce_nat)
+ by (rule sum.cong) (simp_all add: choose_reduce_nat)
also have "\<dots> = (\<Sum>k = 0..Suc n. Suc n - k choose k) +
(\<Sum>k = 0..Suc n. if k=0 then 0 else (Suc n - k choose (k - 1)))"
- by (simp add: setsum.distrib)
+ by (simp add: sum.distrib)
also have "\<dots> = (\<Sum>k = 0..Suc n. Suc n - k choose k) +
(\<Sum>j = 0..n. n - j choose j)"
- by (metis setsum_choose_drop_zero)
+ by (metis sum_choose_drop_zero)
finally show ?case using 3
by simp
qed
--- a/src/HOL/Old_Number_Theory/Euler.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Euler.thy Mon Oct 17 14:37:32 2016 +0200
@@ -108,8 +108,8 @@
lemma Union_SetS_finite: "2 < p ==> finite (\<Union>(SetS a p))"
by (auto simp add: SetS_finite SetS_elems_finite)
-lemma card_setsum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set);
- \<forall>X \<in> S. card X = n |] ==> setsum card S = setsum (%x. n) S"
+lemma card_sum_aux: "[| finite S; \<forall>X \<in> S. finite (X::int set);
+ \<forall>X \<in> S. card X = n |] ==> sum card S = sum (%x. n) S"
by (induct set: finite) auto
lemma SetS_card:
@@ -120,14 +120,14 @@
proof -
have "p - 1 = int(card(\<Union>(SetS a p)))"
by (auto simp add: assms MultInvPair_prop2 SRStar_card)
- also have "... = int (setsum card (SetS a p))"
+ also have "... = int (sum card (SetS a p))"
by (auto simp add: assms SetS_finite SetS_elems_finite
MultInvPair_prop1c [of p a] card_Union_disjoint)
- also have "... = int(setsum (%x.2) (SetS a p))"
+ also have "... = int(sum (%x.2) (SetS a p))"
using assms by (auto simp add: SetS_elems_card SetS_finite SetS_elems_finite
- card_setsum_aux simp del: setsum_constant)
+ card_sum_aux simp del: sum_constant)
also have "... = 2 * int(card( SetS a p))"
- by (auto simp add: assms SetS_finite setsum_const2)
+ by (auto simp add: assms SetS_finite sum_const2)
finally show ?thesis .
qed
then show ?thesis by auto
--- a/src/HOL/Old_Number_Theory/Fib.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Fib.thy Mon Oct 17 14:37:32 2016 +0200
@@ -138,7 +138,7 @@
apply (simp_all add: gcd_non_0 gcd_commute gcd_fib_mod)
done
-theorem fib_mult_eq_setsum:
+theorem fib_mult_eq_sum:
"fib (Suc n) * fib n = (\<Sum>k \<in> {..n}. fib k * fib k)"
apply (induct n rule: fib.induct)
apply (auto simp add: atMost_Suc fib_2)
--- a/src/HOL/Old_Number_Theory/Finite2.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Finite2.thy Mon Oct 17 14:37:32 2016 +0200
@@ -16,9 +16,9 @@
subsection \<open>Useful properties of sums and products\<close>
-lemma setsum_same_function_zcong:
+lemma sum_same_function_zcong:
assumes a: "\<forall>x \<in> S. [f x = g x](mod m)"
- shows "[setsum f S = setsum g S] (mod m)"
+ shows "[sum f S = sum g S] (mod m)"
proof cases
assume "finite S"
thus ?thesis using a by induct (simp_all add: zcong_zadd)
@@ -36,15 +36,15 @@
assume "infinite S" thus ?thesis by simp
qed
-lemma setsum_const: "finite X ==> setsum (%x. (c :: int)) X = c * int(card X)"
+lemma sum_const: "finite X ==> sum (%x. (c :: int)) X = c * int(card X)"
by (simp add: of_nat_mult)
-lemma setsum_const2: "finite X ==> int (setsum (%x. (c :: nat)) X) =
+lemma sum_const2: "finite X ==> int (sum (%x. (c :: nat)) X) =
int(c) * int(card X)"
by (simp add: of_nat_mult)
-lemma setsum_const_mult: "finite A ==> setsum (%x. c * ((f x)::int)) A =
- c * setsum f A"
+lemma sum_const_mult: "finite A ==> sum (%x. c * ((f x)::int)) A =
+ c * sum f A"
by (induct set: finite) (auto simp add: distrib_left)
--- a/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy Mon Oct 17 14:37:32 2016 +0200
@@ -16,110 +16,110 @@
context GAUSS
begin
-lemma QRLemma1: "a * setsum id A =
- p * setsum (%x. ((x * a) div p)) A + setsum id D + setsum id E"
+lemma QRLemma1: "a * sum id A =
+ p * sum (%x. ((x * a) div p)) A + sum id D + sum id E"
proof -
- from finite_A have "a * setsum id A = setsum (%x. a * x) A"
- by (auto simp add: setsum_const_mult id_def)
- also have "setsum (%x. a * x) = setsum (%x. x * a)"
+ from finite_A have "a * sum id A = sum (%x. a * x) A"
+ by (auto simp add: sum_const_mult id_def)
+ also have "sum (%x. a * x) = sum (%x. x * a)"
by (auto simp add: mult.commute)
- also have "setsum (%x. x * a) A = setsum id B"
- by (simp add: B_def setsum.reindex [OF inj_on_xa_A])
- also have "... = setsum (%x. p * (x div p) + StandardRes p x) B"
+ also have "sum (%x. x * a) A = sum id B"
+ by (simp add: B_def sum.reindex [OF inj_on_xa_A])
+ also have "... = sum (%x. p * (x div p) + StandardRes p x) B"
by (auto simp add: StandardRes_def mult_div_mod_eq [symmetric])
- also have "... = setsum (%x. p * (x div p)) B + setsum (StandardRes p) B"
- by (rule setsum.distrib)
- also have "setsum (StandardRes p) B = setsum id C"
- by (auto simp add: C_def setsum.reindex [OF SR_B_inj])
- also from C_eq have "... = setsum id (D \<union> E)"
+ also have "... = sum (%x. p * (x div p)) B + sum (StandardRes p) B"
+ by (rule sum.distrib)
+ also have "sum (StandardRes p) B = sum id C"
+ by (auto simp add: C_def sum.reindex [OF SR_B_inj])
+ also from C_eq have "... = sum id (D \<union> E)"
by auto
- also from finite_D finite_E have "... = setsum id D + setsum id E"
- by (rule setsum.union_disjoint) (auto simp add: D_def E_def)
- also have "setsum (%x. p * (x div p)) B =
- setsum ((%x. p * (x div p)) o (%x. (x * a))) A"
- by (auto simp add: B_def setsum.reindex inj_on_xa_A)
- also have "... = setsum (%x. p * ((x * a) div p)) A"
+ also from finite_D finite_E have "... = sum id D + sum id E"
+ by (rule sum.union_disjoint) (auto simp add: D_def E_def)
+ also have "sum (%x. p * (x div p)) B =
+ sum ((%x. p * (x div p)) o (%x. (x * a))) A"
+ by (auto simp add: B_def sum.reindex inj_on_xa_A)
+ also have "... = sum (%x. p * ((x * a) div p)) A"
by (auto simp add: o_def)
- also from finite_A have "setsum (%x. p * ((x * a) div p)) A =
- p * setsum (%x. ((x * a) div p)) A"
- by (auto simp add: setsum_const_mult)
+ also from finite_A have "sum (%x. p * ((x * a) div p)) A =
+ p * sum (%x. ((x * a) div p)) A"
+ by (auto simp add: sum_const_mult)
finally show ?thesis by arith
qed
-lemma QRLemma2: "setsum id A = p * int (card E) - setsum id E +
- setsum id D"
+lemma QRLemma2: "sum id A = p * int (card E) - sum id E +
+ sum id D"
proof -
- from F_Un_D_eq_A have "setsum id A = setsum id (D \<union> F)"
+ from F_Un_D_eq_A have "sum id A = sum id (D \<union> F)"
by (simp add: Un_commute)
also from F_D_disj finite_D finite_F
- have "... = setsum id D + setsum id F"
- by (auto simp add: Int_commute intro: setsum.union_disjoint)
+ have "... = sum id D + sum id F"
+ by (auto simp add: Int_commute intro: sum.union_disjoint)
also from F_def have "F = (%x. (p - x)) ` E"
by auto
- also from finite_E inj_on_pminusx_E have "setsum id ((%x. (p - x)) ` E) =
- setsum (%x. (p - x)) E"
- by (auto simp add: setsum.reindex)
- also from finite_E have "setsum (op - p) E = setsum (%x. p) E - setsum id E"
- by (auto simp add: setsum_subtractf id_def)
- also from finite_E have "setsum (%x. p) E = p * int(card E)"
- by (intro setsum_const)
+ also from finite_E inj_on_pminusx_E have "sum id ((%x. (p - x)) ` E) =
+ sum (%x. (p - x)) E"
+ by (auto simp add: sum.reindex)
+ also from finite_E have "sum (op - p) E = sum (%x. p) E - sum id E"
+ by (auto simp add: sum_subtractf id_def)
+ also from finite_E have "sum (%x. p) E = p * int(card E)"
+ by (intro sum_const)
finally show ?thesis
by arith
qed
-lemma QRLemma3: "(a - 1) * setsum id A =
- p * (setsum (%x. ((x * a) div p)) A - int(card E)) + 2 * setsum id E"
+lemma QRLemma3: "(a - 1) * sum id A =
+ p * (sum (%x. ((x * a) div p)) A - int(card E)) + 2 * sum id E"
proof -
- have "(a - 1) * setsum id A = a * setsum id A - setsum id A"
+ have "(a - 1) * sum id A = a * sum id A - sum id A"
by (auto simp add: left_diff_distrib)
also note QRLemma1
- also from QRLemma2 have "p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
- setsum id E - setsum id A =
- p * (\<Sum>x \<in> A. x * a div p) + setsum id D +
- setsum id E - (p * int (card E) - setsum id E + setsum id D)"
+ also from QRLemma2 have "p * (\<Sum>x \<in> A. x * a div p) + sum id D +
+ sum id E - sum id A =
+ p * (\<Sum>x \<in> A. x * a div p) + sum id D +
+ sum id E - (p * int (card E) - sum id E + sum id D)"
by auto
also have "... = p * (\<Sum>x \<in> A. x * a div p) -
- p * int (card E) + 2 * setsum id E"
+ p * int (card E) + 2 * sum id E"
by arith
finally show ?thesis
by (auto simp only: right_diff_distrib)
qed
lemma QRLemma4: "a \<in> zOdd ==>
- (setsum (%x. ((x * a) div p)) A \<in> zEven) = (int(card E): zEven)"
+ (sum (%x. ((x * a) div p)) A \<in> zEven) = (int(card E): zEven)"
proof -
assume a_odd: "a \<in> zOdd"
- from QRLemma3 have a: "p * (setsum (%x. ((x * a) div p)) A - int(card E)) =
- (a - 1) * setsum id A - 2 * setsum id E"
+ from QRLemma3 have a: "p * (sum (%x. ((x * a) div p)) A - int(card E)) =
+ (a - 1) * sum id A - 2 * sum id E"
by arith
from a_odd have "a - 1 \<in> zEven"
by (rule odd_minus_one_even)
- hence "(a - 1) * setsum id A \<in> zEven"
+ hence "(a - 1) * sum id A \<in> zEven"
by (rule even_times_either)
- moreover have "2 * setsum id E \<in> zEven"
+ moreover have "2 * sum id E \<in> zEven"
by (auto simp add: zEven_def)
- ultimately have "(a - 1) * setsum id A - 2 * setsum id E \<in> zEven"
+ ultimately have "(a - 1) * sum id A - 2 * sum id E \<in> zEven"
by (rule even_minus_even)
- with a have "p * (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+ with a have "p * (sum (%x. ((x * a) div p)) A - int(card E)): zEven"
by simp
- hence "p \<in> zEven | (setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+ hence "p \<in> zEven | (sum (%x. ((x * a) div p)) A - int(card E)): zEven"
by (rule EvenOdd.even_product)
- with p_odd have "(setsum (%x. ((x * a) div p)) A - int(card E)): zEven"
+ with p_odd have "(sum (%x. ((x * a) div p)) A - int(card E)): zEven"
by (auto simp add: odd_iff_not_even)
thus ?thesis
by (auto simp only: even_diff [symmetric])
qed
lemma QRLemma5: "a \<in> zOdd ==>
- (-1::int)^(card E) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
+ (-1::int)^(card E) = (-1::int)^(nat(sum (%x. ((x * a) div p)) A))"
proof -
assume "a \<in> zOdd"
from QRLemma4 [OF this] have
- "(int(card E): zEven) = (setsum (%x. ((x * a) div p)) A \<in> zEven)" ..
+ "(int(card E): zEven) = (sum (%x. ((x * a) div p)) A \<in> zEven)" ..
moreover have "0 \<le> int(card E)"
by auto
- moreover have "0 \<le> setsum (%x. ((x * a) div p)) A"
- proof (intro setsum_nonneg)
+ moreover have "0 \<le> sum (%x. ((x * a) div p)) A"
+ proof (intro sum_nonneg)
show "\<forall>x \<in> A. 0 \<le> x * a div p"
proof
fix x
@@ -144,7 +144,7 @@
lemma MainQRLemma: "[| a \<in> zOdd; 0 < a; ~([a = 0] (mod p)); zprime p; 2 < p;
A = {x. 0 < x & x \<le> (p - 1) div 2} |] ==>
- (Legendre a p) = (-1::int)^(nat(setsum (%x. ((x * a) div p)) A))"
+ (Legendre a p) = (-1::int)^(nat(sum (%x. ((x * a) div p)) A))"
apply (subst GAUSS.gauss_lemma)
apply (auto simp add: GAUSS_def)
apply (subst GAUSS.QRLemma5)
@@ -508,7 +508,7 @@
finally show "int (card (f2 j)) = p * j div q" .
qed
-lemma S1_card: "int (card(S1)) = setsum (%j. (q * j) div p) P_set"
+lemma S1_card: "int (card(S1)) = sum (%j. (q * j) div p) P_set"
proof -
have "\<forall>x \<in> P_set. finite (f1 x)"
proof
@@ -520,18 +520,18 @@
by (auto simp add: f1_def)
moreover note P_set_finite
ultimately have "int(card (UNION P_set f1)) =
- setsum (%x. int(card (f1 x))) P_set"
- by(simp add:card_UN_disjoint int_setsum o_def)
+ sum (%x. int(card (f1 x))) P_set"
+ by(simp add:card_UN_disjoint int_sum o_def)
moreover have "S1 = UNION P_set f1"
by (auto simp add: f1_def S_def S1_def S2_def P_set_def Q_set_def aux1a)
- ultimately have "int(card (S1)) = setsum (%j. int(card (f1 j))) P_set"
+ ultimately have "int(card (S1)) = sum (%j. int(card (f1 j))) P_set"
by auto
- also have "... = setsum (%j. q * j div p) P_set"
- using aux3a by(fastforce intro: setsum.cong)
+ also have "... = sum (%j. q * j div p) P_set"
+ using aux3a by(fastforce intro: sum.cong)
finally show ?thesis .
qed
-lemma S2_card: "int (card(S2)) = setsum (%j. (p * j) div q) Q_set"
+lemma S2_card: "int (card(S2)) = sum (%j. (p * j) div q) Q_set"
proof -
have "\<forall>x \<in> Q_set. finite (f2 x)"
proof
@@ -544,30 +544,30 @@
by (auto simp add: f2_def)
moreover note Q_set_finite
ultimately have "int(card (UNION Q_set f2)) =
- setsum (%x. int(card (f2 x))) Q_set"
- by(simp add:card_UN_disjoint int_setsum o_def)
+ sum (%x. int(card (f2 x))) Q_set"
+ by(simp add:card_UN_disjoint int_sum o_def)
moreover have "S2 = UNION Q_set f2"
by (auto simp add: f2_def S_def S1_def S2_def P_set_def Q_set_def aux1b)
- ultimately have "int(card (S2)) = setsum (%j. int(card (f2 j))) Q_set"
+ ultimately have "int(card (S2)) = sum (%j. int(card (f2 j))) Q_set"
by auto
- also have "... = setsum (%j. p * j div q) Q_set"
- using aux3b by(fastforce intro: setsum.cong)
+ also have "... = sum (%j. p * j div q) Q_set"
+ using aux3b by(fastforce intro: sum.cong)
finally show ?thesis .
qed
lemma S1_carda: "int (card(S1)) =
- setsum (%j. (j * q) div p) P_set"
+ sum (%j. (j * q) div p) P_set"
by (auto simp add: S1_card ac_simps)
lemma S2_carda: "int (card(S2)) =
- setsum (%j. (j * p) div q) Q_set"
+ sum (%j. (j * p) div q) Q_set"
by (auto simp add: S2_card ac_simps)
-lemma pq_sum_prop: "(setsum (%j. (j * p) div q) Q_set) +
- (setsum (%j. (j * q) div p) P_set) = ((p - 1) div 2) * ((q - 1) div 2)"
+lemma pq_sum_prop: "(sum (%j. (j * p) div q) Q_set) +
+ (sum (%j. (j * q) div p) P_set) = ((p - 1) div 2) * ((q - 1) div 2)"
proof -
- have "(setsum (%j. (j * p) div q) Q_set) +
- (setsum (%j. (j * q) div p) P_set) = int (card S2) + int (card S1)"
+ have "(sum (%j. (j * p) div q) Q_set) +
+ (sum (%j. (j * q) div p) P_set) = int (card S2) + int (card S1)"
by (auto simp add: S1_carda S2_carda)
also have "... = int (card S1) + int (card S2)"
by auto
@@ -591,7 +591,7 @@
from QRTEMP_axioms have "~([p = 0] (mod q))"
by (auto simp add: pq_prime_neq QRTEMP_def)
with QRTEMP_axioms Q_set_def have a1: "(Legendre p q) = (-1::int) ^
- nat(setsum (%x. ((x * p) div q)) Q_set)"
+ nat(sum (%x. ((x * p) div q)) Q_set)"
apply (rule_tac p = q in MainQRLemma)
apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
done
@@ -600,22 +600,22 @@
apply (simp add: QRTEMP_def)+
done
with QRTEMP_axioms P_set_def have a2: "(Legendre q p) =
- (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
+ (-1::int) ^ nat(sum (%x. ((x * q) div p)) P_set)"
apply (rule_tac p = p in MainQRLemma)
apply (auto simp add: zprime_zOdd_eq_grt_2 QRTEMP_def)
done
from a1 a2 have "(Legendre p q) * (Legendre q p) =
- (-1::int) ^ nat(setsum (%x. ((x * p) div q)) Q_set) *
- (-1::int) ^ nat(setsum (%x. ((x * q) div p)) P_set)"
+ (-1::int) ^ nat(sum (%x. ((x * p) div q)) Q_set) *
+ (-1::int) ^ nat(sum (%x. ((x * q) div p)) P_set)"
by auto
- also have "... = (-1::int) ^ (nat(setsum (%x. ((x * p) div q)) Q_set) +
- nat(setsum (%x. ((x * q) div p)) P_set))"
+ also have "... = (-1::int) ^ (nat(sum (%x. ((x * p) div q)) Q_set) +
+ nat(sum (%x. ((x * q) div p)) P_set))"
by (auto simp add: power_add)
- also have "nat(setsum (%x. ((x * p) div q)) Q_set) +
- nat(setsum (%x. ((x * q) div p)) P_set) =
- nat((setsum (%x. ((x * p) div q)) Q_set) +
- (setsum (%x. ((x * q) div p)) P_set))"
- apply (rule_tac z = "setsum (%x. ((x * p) div q)) Q_set" in
+ also have "nat(sum (%x. ((x * p) div q)) Q_set) +
+ nat(sum (%x. ((x * q) div p)) P_set) =
+ nat((sum (%x. ((x * p) div q)) Q_set) +
+ (sum (%x. ((x * q) div p)) P_set))"
+ apply (rule_tac z = "sum (%x. ((x * p) div q)) Q_set" in
nat_add_distrib [symmetric])
apply (auto simp add: S1_carda [symmetric] S2_carda [symmetric])
done
--- a/src/HOL/Old_Number_Theory/Residues.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Old_Number_Theory/Residues.thy Mon Oct 17 14:37:32 2016 +0200
@@ -129,7 +129,7 @@
done
lemma StandardRes_Sum: "[| finite X; 0 < m |]
- ==> [setsum f X = setsum (StandardRes m o f) X](mod m)"
+ ==> [sum f X = sum (StandardRes m o f) X](mod m)"
apply (rule_tac F = X in finite_induct)
apply (auto intro!: zcong_zadd simp add: StandardRes_prop1)
done
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Quickcheck_Examples.thy Mon Oct 17 14:37:32 2016 +0200
@@ -313,7 +313,7 @@
(is "(\<Sum>x\<leftarrow>_. ?f x) = _")
unfolding sparsify_def scalar_product_def
using assms sum_list_map_filter[where f="?f" and P="\<lambda> i. snd i \<noteq> (0::'a)"]
- by (simp add: sum_list_setsum)
+ by (simp add: sum_list_sum)
*)
definition [simp]: "unzip w = (map fst w, map snd w)"
--- a/src/HOL/Probability/Central_Limit_Theorem.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Central_Limit_Theorem.thy Mon Oct 17 14:37:32 2016 +0200
@@ -54,10 +54,10 @@
by (subst X_distrib [symmetric]) (auto simp: integral_distr)
have "\<phi> n t = char (distr M borel (\<lambda>x. \<Sum>i<n. X i x / sqrt (\<sigma>\<^sup>2 * real n))) t"
- by (auto simp: \<phi>_def S_def setsum_divide_distrib ac_simps)
+ by (auto simp: \<phi>_def S_def sum_divide_distrib ac_simps)
also have "\<dots> = (\<Prod> i < n. \<psi>' n i t)"
unfolding \<psi>'_def
- apply (rule char_distr_setsum)
+ apply (rule char_distr_sum)
apply (rule indep_vars_compose2[where X=X])
apply (rule indep_vars_subset)
apply (rule X_indep)
--- a/src/HOL/Probability/Characteristic_Functions.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Characteristic_Functions.thy Mon Oct 17 14:37:32 2016 +0200
@@ -96,7 +96,7 @@
subsection \<open>Independence\<close>
(* the automation can probably be improved *)
-lemma (in prob_space) char_distr_sum:
+lemma (in prob_space) char_distr_add:
fixes X1 X2 :: "'a \<Rightarrow> real" and t :: real
assumes "indep_var borel X1 borel X2"
shows "char (distr M borel (\<lambda>\<omega>. X1 \<omega> + X2 \<omega>)) t =
@@ -117,12 +117,12 @@
finally show ?thesis .
qed
-lemma (in prob_space) char_distr_setsum:
+lemma (in prob_space) char_distr_sum:
"indep_vars (\<lambda>i. borel) X A \<Longrightarrow>
char (distr M borel (\<lambda>\<omega>. \<Sum>i\<in>A. X i \<omega>)) t = (\<Prod>i\<in>A. char (distr M borel (X i)) t)"
proof (induct A rule: infinite_finite_induct)
case (insert x F) with indep_vars_subset[of "\<lambda>_. borel" X "insert x F" F] show ?case
- by (auto simp add: char_distr_sum indep_vars_setsum)
+ by (auto simp add: char_distr_add indep_vars_sum)
qed (simp_all add: char_def integral_distr prob_space del: distr_const)
subsection \<open>Approximations to $e^{ix}$\<close>
@@ -177,7 +177,7 @@
unfolding of_nat_mult[symmetric]
by (simp add: complex_eq_iff ** of_nat_add[symmetric] del: of_nat_mult of_nat_fact of_nat_add)
show "?P (Suc n)"
- unfolding setsum_atMost_Suc ih f_def CLBINT_I0c_power_mirror_iexp[of _ n]
+ unfolding sum_atMost_Suc ih f_def CLBINT_I0c_power_mirror_iexp[of _ n]
by (simp add: divide_simps add_eq_0_iff *) (simp add: field_simps)
qed
@@ -318,11 +318,11 @@
also have "\<dots> = norm ((CLINT x | M. iexp (t * x) - (\<Sum>k \<le> n. c k x)))"
unfolding char_def by (subst Bochner_Integration.integral_diff[OF integ_iexp]) (auto intro!: integ_c)
also have "\<dots> \<le> expectation (\<lambda>x. cmod (iexp (t * x) - (\<Sum>k \<le> n. c k x)))"
- by (intro integral_norm_bound Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_setsum integ_c) simp
+ by (intro integral_norm_bound Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp
also have "\<dots> \<le> expectation (\<lambda>x. 2 * \<bar>t\<bar> ^ n / fact n * \<bar>x\<bar> ^ n)"
proof (rule integral_mono)
show "integrable M (\<lambda>x. cmod (iexp (t * x) - (\<Sum>k\<le>n. c k x)))"
- by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_setsum integ_c) simp
+ by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp
show "integrable M (\<lambda>x. 2 * \<bar>t\<bar> ^ n / fact n * \<bar>x\<bar> ^ n)"
unfolding power_abs[symmetric]
by (intro integrable_mult_right integrable_abs integrable_moments) simp
@@ -362,12 +362,12 @@
also have "\<dots> = norm ((CLINT x | M. iexp (t * x) - (\<Sum>k \<le> n. c k x)))"
unfolding char_def by (subst Bochner_Integration.integral_diff[OF integ_iexp]) (auto intro!: integ_c)
also have "\<dots> \<le> expectation (\<lambda>x. cmod (iexp (t * x) - (\<Sum>k \<le> n. c k x)))"
- by (intro integral_norm_bound Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_setsum integ_c) simp
+ by (intro integral_norm_bound Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp
also have "\<dots> \<le> expectation (\<lambda>x. min (2 * \<bar>t * x\<bar>^n / fact n) (\<bar>t * x\<bar>^(Suc n) / fact (Suc n)))"
(is "_ \<le> expectation ?f")
proof (rule integral_mono)
show "integrable M (\<lambda>x. cmod (iexp (t * x) - (\<Sum>k\<le>n. c k x)))"
- by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_setsum integ_c) simp
+ by (intro integrable_norm Bochner_Integration.integrable_diff integ_iexp Bochner_Integration.integrable_sum integ_c) simp
show "integrable M ?f"
by (rule Bochner_Integration.integrable_bound[where f="\<lambda>x. 2 * \<bar>t * x\<bar> ^ n / fact n"])
(auto simp: integrable_moments power_abs[symmetric] power_mult_distrib)
@@ -502,8 +502,8 @@
have "(\<i> * complex_of_real t) ^ (2 * a) / (2 ^ a * fact a) = (- ((complex_of_real t)\<^sup>2 / 2)) ^ a / fact a" for a
by (subst power_mult) (simp add: field_simps uminus_power_if power_mult)
then have *: "?f (2 * n) = complex_of_real (\<Sum>k < Suc n. (1 / fact k) * (- (t^2) / 2)^k)" for n :: nat
- unfolding of_real_setsum
- by (intro setsum.reindex_bij_witness_not_neutral[symmetric, where
+ unfolding of_real_sum
+ by (intro sum.reindex_bij_witness_not_neutral[symmetric, where
i="\<lambda>n. n div 2" and j="\<lambda>n. 2 * n" and T'="{i. i \<le> 2 * n \<and> odd i}" and S'="{}"])
(auto simp: integral_std_normal_distribution_moment_odd std_normal_distribution_even_moments)
show "(\<lambda>n. ?f (2 * n)) \<longlonglongrightarrow> exp (-(t^2) / 2)"
--- a/src/HOL/Probability/Distributions.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Distributions.thy Mon Oct 17 14:37:32 2016 +0200
@@ -102,7 +102,7 @@
qed
qed auto
also have "\<dots> = ennreal (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n))"
- by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] setsum.If_cases)
+ by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] sum.If_cases)
also have "\<dots> = ennreal ((1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n)) * fact k) * ennreal (inverse (fact k))"
by (subst ennreal_mult''[symmetric]) (auto intro!: arg_cong[where f=ennreal])
finally show ?thesis
@@ -149,14 +149,14 @@
assume "0 \<le> x"
have "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) =
exp (- (l * x)) * (\<Sum>n\<le>k. (l * x) ^ n / fact n)"
- unfolding setsum_distrib_left by (intro setsum.cong) (auto simp: field_simps)
+ unfolding sum_distrib_left by (intro sum.cong) (auto simp: field_simps)
also have "\<dots> = (\<Sum>n\<le>k. (l * x) ^ n / fact n) / exp (l * x)"
by (simp add: exp_minus field_simps)
also have "\<dots> \<le> 1"
proof (subst divide_le_eq_1_pos)
show "(\<Sum>n\<le>k. (l * x) ^ n / fact n) \<le> exp (l * x)"
using \<open>0 < l\<close> \<open>0 \<le> x\<close> summable_exp_generic[of "l * x"]
- by (auto simp: exp_def divide_inverse ac_simps intro!: setsum_le_suminf)
+ by (auto simp: exp_def divide_inverse ac_simps intro!: sum_le_suminf)
qed simp
finally show "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) \<le> 1" .
qed
@@ -457,7 +457,7 @@
case (insert i I)
then have "distributed M lborel (\<lambda>x. min (X i x) (Min ((\<lambda>i. X i x)`I))) (exponential_density (l i + (\<Sum>i\<in>I. l i)))"
by (intro exponential_distributed_min indep_vars_Min insert)
- (auto intro: indep_vars_subset setsum_pos)
+ (auto intro: indep_vars_subset sum_pos)
then show ?case
using insert by simp
qed
@@ -557,7 +557,7 @@
apply auto
done
-lemma (in prob_space) erlang_distributed_setsum:
+lemma (in prob_space) erlang_distributed_sum:
assumes finI : "finite I"
assumes A: "I \<noteq> {}"
assumes [simp, arith]: "0 < l"
@@ -570,7 +570,7 @@
next
case (insert i I)
then have "distributed M lborel (\<lambda>x. (X i x) + (\<Sum>i\<in> I. X i x)) (erlang_density (Suc (k i) + Suc ((\<Sum>i\<in>I. Suc (k i)) - 1) - 1) l)"
- by(intro sum_indep_erlang indep_vars_setsum) (auto intro!: indep_vars_subset)
+ by(intro sum_indep_erlang indep_vars_sum) (auto intro!: indep_vars_subset)
also have "(\<lambda>x. (X i x) + (\<Sum>i\<in> I. X i x)) = (\<lambda>x. \<Sum>i\<in>insert i I. X i x)"
using insert by auto
also have "Suc(k i) + Suc ((\<Sum>i\<in>I. Suc (k i)) - 1) - 1 = (\<Sum>i\<in>insert i I. Suc (k i)) - 1"
@@ -578,14 +578,14 @@
finally show ?case by fast
qed
-lemma (in prob_space) exponential_distributed_setsum:
+lemma (in prob_space) exponential_distributed_sum:
assumes finI: "finite I"
assumes A: "I \<noteq> {}"
assumes l: "0 < l"
assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (exponential_density l)"
assumes ind: "indep_vars (\<lambda>i. borel) X I"
shows "distributed M lborel (\<lambda>x. \<Sum>i\<in>I. X i x) (erlang_density ((card I) - 1) l)"
- using erlang_distributed_setsum[OF assms] by simp
+ using erlang_distributed_sum[OF assms] by simp
lemma (in information_space) entropy_exponential:
assumes l[simp, arith]: "0 < l"
@@ -1248,7 +1248,7 @@
(normal_density 0 (sqrt 2))"
by (subst conv_normal_density_zero_mean) simp_all
-lemma (in prob_space) sum_indep_normal:
+lemma (in prob_space) add_indep_normal:
assumes indep: "indep_var borel X borel Y"
assumes pos_var[arith]: "0 < \<sigma>" "0 < \<tau>"
assumes normalX[simp]: "distributed M lborel X (normal_density \<mu> \<sigma>)"
@@ -1293,14 +1293,14 @@
then have [simp]:"distributed M lborel (\<lambda>x. - Y x) (\<lambda>x. ennreal (normal_density (- \<nu>) \<tau> x))" by simp
have "distributed M lborel (\<lambda>x. X x + (- Y x)) (normal_density (\<mu> + - \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"
- apply (rule sum_indep_normal)
+ apply (rule add_indep_normal)
apply (rule indep_var_compose[unfolded comp_def, of borel _ borel _ "\<lambda>x. x" _ "\<lambda>x. - x"])
apply simp_all
done
then show ?thesis by simp
qed
-lemma (in prob_space) setsum_indep_normal:
+lemma (in prob_space) sum_indep_normal:
assumes "finite I" "I \<noteq> {}" "indep_vars (\<lambda>i. borel) X I"
assumes "\<And>i. i \<in> I \<Longrightarrow> 0 < \<sigma> i"
assumes normal: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (normal_density (\<mu> i) (\<sigma> i))"
@@ -1311,17 +1311,17 @@
next
case (insert i I)
then have 1: "distributed M lborel (\<lambda>x. (X i x) + (\<Sum>i\<in>I. X i x))
- (normal_density (\<mu> i + setsum \<mu> I) (sqrt ((\<sigma> i)\<^sup>2 + (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2))\<^sup>2)))"
- apply (intro sum_indep_normal indep_vars_setsum insert real_sqrt_gt_zero)
- apply (auto intro: indep_vars_subset intro!: setsum_pos)
+ (normal_density (\<mu> i + sum \<mu> I) (sqrt ((\<sigma> i)\<^sup>2 + (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2))\<^sup>2)))"
+ apply (intro add_indep_normal indep_vars_sum insert real_sqrt_gt_zero)
+ apply (auto intro: indep_vars_subset intro!: sum_pos)
apply fastforce
done
have 2: "(\<lambda>x. (X i x)+ (\<Sum>i\<in>I. X i x)) = (\<lambda>x. (\<Sum>j\<in>insert i I. X j x))"
- "\<mu> i + setsum \<mu> I = setsum \<mu> (insert i I)"
+ "\<mu> i + sum \<mu> I = sum \<mu> (insert i I)"
using insert by auto
have 3: "(sqrt ((\<sigma> i)\<^sup>2 + (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2))\<^sup>2)) = (sqrt (\<Sum>i\<in>insert i I. (\<sigma> i)\<^sup>2))"
- using insert by (simp add: setsum_nonneg)
+ using insert by (simp add: sum_nonneg)
show ?case using 1 2 3 by simp
qed
--- a/src/HOL/Probability/Independent_Family.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Independent_Family.thy Mon Oct 17 14:37:32 2016 +0200
@@ -991,7 +991,7 @@
unfolding indep_var_def .
qed
-lemma (in prob_space) indep_vars_setsum:
+lemma (in prob_space) indep_vars_sum:
fixes X :: "'i \<Rightarrow> 'a \<Rightarrow> real"
assumes I: "finite I" "i \<notin> I" and indep: "indep_vars (\<lambda>_. borel) X (insert i I)"
shows "indep_var borel (X i) borel (\<lambda>\<omega>. \<Sum>i\<in>I. X i \<omega>)"
--- a/src/HOL/Probability/Information.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Information.thy Mon Oct 17 14:37:32 2016 +0200
@@ -16,9 +16,9 @@
lemma log_less: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x < y \<Longrightarrow> log a x < log a y"
by (subst log_less_cancel_iff) auto
-lemma setsum_cartesian_product':
- "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
- unfolding setsum.cartesian_product by simp
+lemma sum_cartesian_product':
+ "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. sum (\<lambda>y. f (x, y)) B)"
+ unfolding sum.cartesian_product by simp
lemma split_pairs:
"((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
@@ -747,8 +747,8 @@
by auto
with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M))) =
(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
- by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum.If_cases split_beta'
- intro!: setsum.cong)
+ by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite sum.If_cases split_beta'
+ intro!: sum.cong)
qed (insert X Y XY, auto simp: simple_distributed_def)
lemma (in information_space)
@@ -760,7 +760,7 @@
proof (subst mutual_information_simple_distributed[OF Px Py Pxy])
have "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) =
(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. 0)"
- by (intro setsum.cong) (auto simp: ae)
+ by (intro sum.cong) (auto simp: ae)
then show "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M.
Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp
qed
@@ -1479,7 +1479,7 @@
have "(\<lambda>(x, y, z). ?i x y z) = (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)"
by (auto intro!: ext)
then show "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x)) ` space M. ?j x y z)"
- by (auto intro!: setsum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite setsum.If_cases split_beta')
+ by (auto intro!: sum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite sum.If_cases split_beta')
qed (insert Pz Pyz Pxz Pxyz, auto intro: measure_nonneg)
lemma (in information_space) conditional_mutual_information_nonneg:
@@ -1659,7 +1659,7 @@
by auto
from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
- by (auto intro!: setsum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite eq setsum.If_cases split_beta')
+ by (auto intro!: sum.cong simp add: \<open>?P = ?C\<close> lebesgue_integral_count_space_finite simple_distributed_finite eq sum.If_cases split_beta')
qed (insert Y XY, auto)
lemma (in information_space) conditional_mutual_information_eq_conditional_entropy:
@@ -1697,7 +1697,7 @@
then show ?thesis
apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy])
apply (subst conditional_entropy_eq[OF Py Pxy])
- apply (auto intro!: setsum.cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum.reindex[OF inj]
+ apply (auto intro!: sum.cong simp: Pxxy_eq sum_negf[symmetric] eq sum.reindex[OF inj]
log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space)
using Py[THEN simple_distributed] Pxy[THEN simple_distributed]
apply (auto simp add: not_le AE_count_space less_le antisym
@@ -1902,13 +1902,13 @@
have "\<H>(\<lambda>x. (X x, Y x)) = - (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` space M. ?f x * log b (?f x))"
using XY by (rule entropy_simple_distributed)
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>(x, y). (y, x)) ` (\<lambda>x. (X x, Y x)) ` space M. ?g x * log b (?g x))"
- by (subst (2) setsum.reindex) (auto simp: inj_on_def intro!: setsum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
+ by (subst (2) sum.reindex) (auto simp: inj_on_def intro!: sum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))"
- by (auto intro!: setsum.cong)
+ by (auto intro!: sum.cong)
also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
by (subst entropy_distr[OF simple_distributed_joint[OF YX]])
(auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
- cong del: setsum.strong_cong intro!: setsum.mono_neutral_left measure_nonneg)
+ cong del: sum.strong_cong intro!: sum.mono_neutral_left measure_nonneg)
finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
then show ?thesis
unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp
@@ -1927,8 +1927,8 @@
apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]])
apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
unfolding eq
- apply (subst setsum.reindex[OF inj])
- apply (auto intro!: setsum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
+ apply (subst sum.reindex[OF inj])
+ apply (auto intro!: sum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
done
qed
@@ -1953,7 +1953,7 @@
unfolding o_assoc
apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]])
apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\<lambda>x. prob (X -` {x} \<inter> space M)"])
- apply (auto intro!: setsum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def measure_nonneg)
+ apply (auto intro!: sum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def measure_nonneg)
done
also have "... \<le> \<H>(f \<circ> X)"
using entropy_data_processing[OF sf] .
--- a/src/HOL/Probability/PMF_Impl.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/PMF_Impl.thy Mon Oct 17 14:37:32 2016 +0200
@@ -27,7 +27,7 @@
by (subst nn_integral_count_space'[of "Mapping.keys m"])
(auto simp: Mapping.lookup_default_def keys_is_none_rep Option.is_none_def)
also from assms have "\<dots> = ennreal (\<Sum>k\<in>Mapping.keys m. Mapping.lookup_default 0 m k)"
- by (intro setsum_ennreal)
+ by (intro sum_ennreal)
(auto simp: Mapping.lookup_default_def All_mapping_def split: option.splits)
finally show ?thesis .
qed
@@ -317,12 +317,12 @@
define prob where "prob = (\<Sum>x\<in>C. pmf p x)"
also note C_def
also from assms have "(\<Sum>x\<in>A \<inter> set_pmf p. pmf p x) = (\<Sum>x\<in>A. pmf p x)"
- by (intro setsum.mono_neutral_left) (auto simp: set_pmf_eq)
+ by (intro sum.mono_neutral_left) (auto simp: set_pmf_eq)
finally have prob1: "prob = (\<Sum>x\<in>A. pmf p x)" .
hence prob2: "prob = measure_pmf.prob p A"
using assms by (subst measure_measure_pmf_finite) simp_all
have prob3: "prob = 0 \<longleftrightarrow> A \<inter> set_pmf p = {}"
- by (subst prob1, subst setsum_nonneg_eq_0_iff) (auto simp: set_pmf_eq assms)
+ by (subst prob1, subst sum_nonneg_eq_0_iff) (auto simp: set_pmf_eq assms)
from assms have prob4: "prob = measure_pmf.prob p C"
unfolding prob_def by (intro measure_measure_pmf_finite [symmetric]) (simp_all add: C_def)
@@ -571,4 +571,4 @@
end
-end
\ No newline at end of file
+end
--- a/src/HOL/Probability/Probability_Mass_Function.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Probability_Mass_Function.thy Mon Oct 17 14:37:32 2016 +0200
@@ -205,13 +205,13 @@
using emeasure_pmf_single[of M x] by(simp add: measure_pmf.emeasure_eq_measure pmf_nonneg measure_nonneg)
lemma emeasure_measure_pmf_finite: "finite S \<Longrightarrow> emeasure (measure_pmf M) S = (\<Sum>s\<in>S. pmf M s)"
- by (subst emeasure_eq_setsum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)
+ by (subst emeasure_eq_sum_singleton) (auto simp: emeasure_pmf_single pmf_nonneg)
-lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = setsum (pmf M) S"
+lemma measure_measure_pmf_finite: "finite S \<Longrightarrow> measure (measure_pmf M) S = sum (pmf M) S"
using emeasure_measure_pmf_finite[of S M]
- by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg setsum_nonneg pmf_nonneg)
+ by (simp add: measure_pmf.emeasure_eq_measure measure_nonneg sum_nonneg pmf_nonneg)
-lemma setsum_pmf_eq_1:
+lemma sum_pmf_eq_1:
assumes "finite A" "set_pmf p \<subseteq> A"
shows "(\<Sum>x\<in>A. pmf p x) = 1"
proof -
@@ -575,7 +575,7 @@
apply (subst integral_measure_pmf_real[where A="{b}"])
apply (auto simp: indicator_eq_0_iff)
apply (subst integral_measure_pmf_real[where A="{a}"])
- apply (auto simp: indicator_eq_0_iff setsum_nonneg_eq_0_iff pmf_nonneg)
+ apply (auto simp: indicator_eq_0_iff sum_nonneg_eq_0_iff pmf_nonneg)
done
lemma set_pair_pmf[simp]: "set_pmf (pair_pmf A B) = set_pmf A \<times> set_pmf B"
@@ -752,7 +752,7 @@
unfolding measure_pmf_eq_density
apply (simp add: integral_density)
apply (subst lebesgue_integral_count_space_finite_support)
- apply (auto intro!: finite_subset[OF _ \<open>finite A\<close>] setsum.mono_neutral_left simp: pmf_eq_0_set_pmf)
+ apply (auto intro!: finite_subset[OF _ \<open>finite A\<close>] sum.mono_neutral_left simp: pmf_eq_0_set_pmf)
done
lemma continuous_on_LINT_pmf: -- \<open>This is dominated convergence!?\<close>
@@ -764,7 +764,7 @@
assume "finite M" with f show ?thesis
using integral_measure_pmf[OF \<open>finite M\<close>]
by (subst integral_measure_pmf[OF \<open>finite M\<close>])
- (auto intro!: continuous_on_setsum continuous_on_scaleR continuous_on_const)
+ (auto intro!: continuous_on_sum continuous_on_scaleR continuous_on_const)
next
assume "infinite M"
let ?f = "\<lambda>i x. pmf (map_pmf (to_nat_on M) M) i *\<^sub>R f (from_nat_into M i) x"
@@ -772,7 +772,7 @@
show ?thesis
proof (rule uniform_limit_theorem)
show "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>a. \<Sum>i<n. ?f i a)"
- by (intro always_eventually allI continuous_on_setsum continuous_on_scaleR continuous_on_const f
+ by (intro always_eventually allI continuous_on_sum continuous_on_scaleR continuous_on_const f
from_nat_into set_pmf_not_empty)
show "uniform_limit A (\<lambda>n a. \<Sum>i<n. ?f i a) (\<lambda>a. LINT i|M. f i a) sequentially"
proof (subst uniform_limit_cong[where g="\<lambda>n a. \<Sum>i<n. ?f i a"])
@@ -1677,8 +1677,8 @@
show "(\<integral>\<^sup>+ x. ennreal (real (count M x) / real (size M)) \<partial>count_space UNIV) = 1"
using M_not_empty
by (simp add: zero_less_divide_iff nn_integral_count_space nonempty_has_size
- setsum_divide_distrib[symmetric])
- (auto simp: size_multiset_overloaded_eq intro!: setsum.cong)
+ sum_divide_distrib[symmetric])
+ (auto simp: size_multiset_overloaded_eq intro!: sum.cong)
qed simp
lemma pmf_of_multiset[simp]: "pmf pmf_of_multiset x = count M x / size M"
@@ -1712,17 +1712,17 @@
lemma emeasure_pmf_of_set_space[simp]: "emeasure pmf_of_set S = 1"
by (rule measure_pmf.emeasure_eq_1_AE) (auto simp: AE_measure_pmf_iff)
-lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = setsum f S / card S"
+lemma nn_integral_pmf_of_set: "nn_integral (measure_pmf pmf_of_set) f = sum f S / card S"
by (subst nn_integral_measure_pmf_finite)
- (simp_all add: setsum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
+ (simp_all add: sum_distrib_right[symmetric] card_gt_0_iff S_not_empty S_finite divide_ennreal_def
divide_ennreal[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_times_divide)
-lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = setsum f S / card S"
- by (subst integral_measure_pmf[of S]) (auto simp: S_finite setsum_divide_distrib)
+lemma integral_pmf_of_set: "integral\<^sup>L (measure_pmf pmf_of_set) f = sum f S / card S"
+ by (subst integral_measure_pmf[of S]) (auto simp: S_finite sum_divide_distrib)
lemma emeasure_pmf_of_set: "emeasure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
by (subst nn_integral_indicator[symmetric], simp)
- (simp add: S_finite S_not_empty card_gt_0_iff indicator_def setsum.If_cases divide_ennreal
+ (simp add: S_finite S_not_empty card_gt_0_iff indicator_def sum.If_cases divide_ennreal
ennreal_of_nat_eq_real_of_nat nn_integral_pmf_of_set)
lemma measure_pmf_of_set: "measure (measure_pmf pmf_of_set) A = card (S \<inter> A) / card S"
@@ -1752,8 +1752,8 @@
with assms have "ennreal ?lhs = ennreal ?rhs"
by (subst ennreal_pmf_bind)
(simp_all add: nn_integral_pmf_of_set max_def pmf_nonneg divide_ennreal [symmetric]
- setsum_nonneg ennreal_of_nat_eq_real_of_nat)
- thus ?thesis by (subst (asm) ennreal_inj) (auto intro!: setsum_nonneg divide_nonneg_nonneg)
+ sum_nonneg ennreal_of_nat_eq_real_of_nat)
+ thus ?thesis by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg divide_nonneg_nonneg)
qed
lemma pmf_of_set_singleton: "pmf_of_set {x} = return_pmf x"
@@ -1801,12 +1801,12 @@
also from assms
have "indicator (\<Union>x\<in>A. f x) x / real \<dots> =
indicator (\<Union>x\<in>A. f x) x / (n * real (card A))"
- by (simp add: setsum_divide_distrib [symmetric] mult_ac)
+ by (simp add: sum_divide_distrib [symmetric] mult_ac)
also from assms have "indicator (\<Union>x\<in>A. f x) x = (\<Sum>y\<in>A. indicator (f y) x)"
by (intro indicator_UN_disjoint) simp_all
also from assms have "ereal ((\<Sum>y\<in>A. indicator (f y) x) / (real n * real (card A))) =
ereal (pmf ?rhs x)"
- by (subst pmf_bind_pmf_of_set) (simp_all add: setsum_divide_distrib)
+ by (subst pmf_bind_pmf_of_set) (simp_all add: sum_divide_distrib)
finally show "pmf ?lhs x = pmf ?rhs x" by simp
qed
@@ -2043,32 +2043,32 @@
by (subst nn_integral_measure_pmf_finite) (simp_all add: finite_set_pmf_of_list pmf_pmf_of_list Int_def)
also from assms
have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A. sum_list (map snd [z\<leftarrow>xs . fst z = x]))"
- by (subst setsum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)
+ by (subst sum_ennreal) (auto simp: pmf_of_list_wf_def intro!: sum_list_nonneg)
also have "\<dots> = ennreal (\<Sum>x\<in>set_pmf (pmf_of_list xs) \<inter> A.
indicator A x * pmf (pmf_of_list xs) x)" (is "_ = ennreal ?S")
- using assms by (intro ennreal_cong setsum.cong) (auto simp: pmf_pmf_of_list)
+ using assms by (intro ennreal_cong sum.cong) (auto simp: pmf_pmf_of_list)
also have "?S = (\<Sum>x\<in>set_pmf (pmf_of_list xs). indicator A x * pmf (pmf_of_list xs) x)"
- using assms by (intro setsum.mono_neutral_left set_pmf_of_list finite_set_pmf_of_list) auto
+ using assms by (intro sum.mono_neutral_left set_pmf_of_list finite_set_pmf_of_list) auto
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x * pmf (pmf_of_list xs) x)"
- using assms by (intro setsum.mono_neutral_left set_pmf_of_list) (auto simp: set_pmf_eq)
+ using assms by (intro sum.mono_neutral_left set_pmf_of_list) (auto simp: set_pmf_eq)
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). indicator A x *
sum_list (map snd (filter (\<lambda>z. fst z = x) xs)))"
using assms by (simp add: pmf_pmf_of_list)
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). sum_list (map snd (filter (\<lambda>z. fst z = x \<and> x \<in> A) xs)))"
- by (intro setsum.cong) (auto simp: indicator_def)
+ by (intro sum.cong) (auto simp: indicator_def)
also have "\<dots> = (\<Sum>x\<in>set (map fst xs). (\<Sum>xa = 0..<length xs.
if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
- by (intro setsum.cong refl, subst sum_list_map_filter', subst sum_list_setsum_nth) simp
+ by (intro sum.cong refl, subst sum_list_map_filter', subst sum_list_sum_nth) simp
also have "\<dots> = (\<Sum>xa = 0..<length xs. (\<Sum>x\<in>set (map fst xs).
if fst (xs ! xa) = x \<and> x \<in> A then snd (xs ! xa) else 0))"
- by (rule setsum.commute)
+ by (rule sum.commute)
also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then
(\<Sum>x\<in>set (map fst xs). if x = fst (xs ! xa) then snd (xs ! xa) else 0) else 0)"
- by (auto intro!: setsum.cong setsum.neutral)
+ by (auto intro!: sum.cong sum.neutral)
also have "\<dots> = (\<Sum>xa = 0..<length xs. if fst (xs ! xa) \<in> A then snd (xs ! xa) else 0)"
- by (intro setsum.cong refl) (simp_all add: setsum.delta)
+ by (intro sum.cong refl) (simp_all add: sum.delta)
also have "\<dots> = sum_list (map snd (filter (\<lambda>x. fst x \<in> A) xs))"
- by (subst sum_list_map_filter', subst sum_list_setsum_nth) simp_all
+ by (subst sum_list_map_filter', subst sum_list_sum_nth) simp_all
finally show ?thesis .
qed
--- a/src/HOL/Probability/Probability_Measure.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy Mon Oct 17 14:37:32 2016 +0200
@@ -383,7 +383,7 @@
by (intro finite_measure_UNION) auto
qed
-lemma (in prob_space) prob_setsum:
+lemma (in prob_space) prob_sum:
assumes [simp, intro]: "finite I"
assumes P: "\<And>n. n \<in> I \<Longrightarrow> {x\<in>space M. P n x} \<in> events"
assumes Q: "{x\<in>space M. Q x} \<in> events"
@@ -1025,14 +1025,14 @@
qed auto
qed
-lemma (in prob_space) simple_distributed_setsum_space:
+lemma (in prob_space) simple_distributed_sum_space:
assumes X: "simple_distributed M X f"
- shows "setsum f (X`space M) = 1"
+ shows "sum f (X`space M) = 1"
proof -
- from X have "setsum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)"
+ from X have "sum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)"
by (subst finite_measure_finite_Union)
(auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD
- intro!: setsum.cong arg_cong[where f="prob"])
+ intro!: sum.cong arg_cong[where f="prob"])
also have "\<dots> = prob (space M)"
by (auto intro!: arg_cong[where f=prob])
finally show ?thesis
@@ -1058,10 +1058,10 @@
using y Px[THEN simple_distributed_finite]
by (auto simp: AE_count_space nn_integral_count_space_finite)
also have "\<dots> = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
- using Pxy by (intro setsum_ennreal) auto
+ using Pxy by (intro sum_ennreal) auto
finally show ?thesis
using simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy]
- by (subst (asm) ennreal_inj) (auto intro!: setsum_nonneg)
+ by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg)
qed
lemma distributedI_real:
--- a/src/HOL/Probability/Projective_Limit.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy Mon Oct 17 14:37:32 2016 +0200
@@ -274,7 +274,7 @@
also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using \<open>Z _ \<in> generator\<close> \<open>Z' _ \<in> generator\<close>
by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto
also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
- proof (rule setsum_mono)
+ proof (rule sum_mono)
fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
unfolding Z'_def Z_def by simp
@@ -290,14 +290,14 @@
finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
qed
also have "\<dots> = ennreal ((\<Sum> i\<in>{1..n}. (2 powr -enn2real i)) * enn2real ?a)"
- using r by (simp add: setsum_distrib_right ennreal_mult[symmetric])
+ using r by (simp add: sum_distrib_right ennreal_mult[symmetric])
also have "\<dots> < ennreal (1 * enn2real ?a)"
proof (intro ennreal_lessI mult_strict_right_mono)
have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
- by (rule setsum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)
+ by (rule sum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)
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 "sum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
+ sum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: sum_diff1)
also have "\<dots> < 1" by (subst geometric_sum) auto
finally show "(\<Sum>i = 1..n. 2 powr - enn2real i) < 1" by simp
qed (auto simp: r enn2real_positive_iff)
--- a/src/HOL/Probability/SPMF.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/SPMF.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1967,10 +1967,10 @@
"measure (measure_spmf (spmf_of_set S)) A = card (S \<inter> A) / card S"
by(auto simp add: measure_spmf_spmf_of_set measure_pmf_of_set)
-lemma nn_integral_spmf_of_set: "nn_integral (measure_spmf (spmf_of_set A)) f = setsum f A / card A"
+lemma nn_integral_spmf_of_set: "nn_integral (measure_spmf (spmf_of_set A)) f = sum f A / card A"
by(cases "finite A")(auto simp add: spmf_of_set_def nn_integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set)
-lemma integral_spmf_of_set: "integral\<^sup>L (measure_spmf (spmf_of_set A)) f = setsum f A / card A"
+lemma integral_spmf_of_set: "integral\<^sup>L (measure_spmf (spmf_of_set A)) f = sum f A / card A"
by(clarsimp simp add: spmf_of_set_def integral_pmf_of_set card_gt_0_iff simp del: spmf_of_pmf_pmf_of_set)
notepad begin \<comment> \<open>@{const pmf_of_set} is not fully parametric.\<close>
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Oct 17 14:37:32 2016 +0200
@@ -83,7 +83,7 @@
lemma zero_notin_Suc_image[simp]: "0 \<notin> Suc ` A"
by auto
-lemma setprod_setsum_distrib_lists:
+lemma setprod_sum_distrib_lists:
fixes P and S :: "'a set" and f :: "'a \<Rightarrow> _::semiring_0" assumes "finite S"
shows "(\<Sum>ms\<in>{ms. set ms \<subseteq> S \<and> length ms = n \<and> (\<forall>i<n. P i (ms!i))}. \<Prod>x<n. f (ms ! x)) =
(\<Prod>i<n. \<Sum>m\<in>{m\<in>S. P i m}. f m)"
@@ -99,8 +99,8 @@
by force+
show ?case unfolding *
using Suc[of "\<lambda>i. P (Suc i)"]
- by (simp add: setsum.reindex split_conv setsum_cartesian_product'
- lessThan_Suc_eq_insert_0 setprod.reindex setsum_distrib_right[symmetric] setsum_distrib_left[symmetric])
+ by (simp add: sum.reindex split_conv sum_cartesian_product'
+ lessThan_Suc_eq_insert_0 setprod.reindex sum_distrib_right[symmetric] sum_distrib_left[symmetric])
qed
declare space_point_measure[simp]
@@ -111,7 +111,7 @@
"finite \<Omega> \<Longrightarrow> A \<subseteq> \<Omega> \<Longrightarrow> (\<And>x. x \<in> \<Omega> \<Longrightarrow> 0 \<le> p x) \<Longrightarrow>
measure (point_measure \<Omega> (\<lambda>x. ennreal (p x))) A = (\<Sum>i\<in>A. p i)"
unfolding measure_def
- by (subst emeasure_point_measure_finite) (auto simp: subset_eq setsum_nonneg)
+ by (subst emeasure_point_measure_finite) (auto simp: subset_eq sum_nonneg)
locale finite_information =
fixes \<Omega> :: "'a set"
@@ -122,8 +122,8 @@
and positive_p[simp, intro]: "\<And>x. 0 \<le> p x"
and b_gt_1[simp, intro]: "1 < b"
-lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
- by (auto intro!: setsum_nonneg)
+lemma (in finite_information) positive_p_sum[simp]: "0 \<le> sum p X"
+ by (auto intro!: sum_nonneg)
sublocale finite_information \<subseteq> prob_space "point_measure \<Omega> p"
by standard (simp add: one_ereal_def emeasure_point_measure_finite)
@@ -131,7 +131,7 @@
sublocale finite_information \<subseteq> information_space "point_measure \<Omega> p" b
by standard simp
-lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> prob A = setsum p A"
+lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> prob A = sum p A"
by (auto simp: measure_point_measure)
locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b
@@ -158,9 +158,9 @@
have [simp]: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (force intro!: inj_onI)
- note setsum_distrib_left[symmetric, simp]
- note setsum_distrib_right[symmetric, simp]
- note setsum_cartesian_product'[simp]
+ note sum_distrib_left[symmetric, simp]
+ note sum_distrib_right[symmetric, simp]
+ note sum_cartesian_product'[simp]
have "(\<Sum>ms | set ms \<subseteq> messages \<and> length ms = n. \<Prod>x<length ms. M (ms ! x)) = 1"
proof (induct n)
@@ -169,9 +169,9 @@
case (Suc n)
then show ?case
by (simp add: lists_length_Suc_eq lessThan_Suc_eq_insert_0
- setsum.reindex setprod.reindex)
+ sum.reindex setprod.reindex)
qed
- then show "setsum P msgs = 1"
+ then show "sum P msgs = 1"
unfolding msgs_def P_def by simp
fix x
have "\<And> A f. 0 \<le> (\<Prod>x\<in>A. M (f x))" by (auto simp: setprod_nonneg)
@@ -211,7 +211,7 @@
qed
abbreviation
- "p A \<equiv> setsum P A"
+ "p A \<equiv> sum P A"
abbreviation
"\<mu> \<equiv> point_measure msgs P"
@@ -254,9 +254,9 @@
show "(\<P>(fst) {k}) = K k"
apply (simp add: \<mu>'_eq)
apply (simp add: * P_def)
- apply (simp add: setsum_cartesian_product')
- using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] \<open>k \<in> keys\<close>
- by (auto simp add: setsum_distrib_left[symmetric] subset_eq setprod.neutral_const)
+ apply (simp add: sum_cartesian_product')
+ using setprod_sum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] \<open>k \<in> keys\<close>
+ by (auto simp add: sum_distrib_left[symmetric] subset_eq setprod.neutral_const)
qed
lemma fst_image_msgs[simp]: "fst`msgs = keys"
@@ -268,7 +268,7 @@
unfolding msgs_def fst_image_times if_not_P[OF *] by simp
qed
-lemma setsum_distribution_cut:
+lemma sum_distribution_cut:
"\<P>(X) {x} = (\<Sum>y \<in> Y`space \<mu>. \<P>(X ; Y) {(x, y)})"
by (subst finite_measure_finite_Union[symmetric])
(auto simp add: disjoint_family_on_def inj_on_def
@@ -298,8 +298,8 @@
show "- (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {x}) * log b (\<P>(X ; Y) {x})) =
- (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` msgs. (\<P>(Y ; X) {x}) * log b (\<P>(Y ; X) {x}))"
unfolding eq
- apply (subst setsum.reindex)
- apply (auto simp: vimage_def inj_on_def intro!: setsum.cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])
+ apply (subst sum.reindex)
+ apply (auto simp: vimage_def inj_on_def intro!: sum.cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])
done
qed simp_all
@@ -315,15 +315,15 @@
proof -
have "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
- (\<Sum>x\<in>X`msgs. (\<Sum>y\<in>Y`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
- unfolding setsum.cartesian_product
- apply (intro arg_cong[where f=uminus] setsum.mono_neutral_left)
+ unfolding sum.cartesian_product
+ apply (intro arg_cong[where f=uminus] sum.mono_neutral_left)
apply (auto simp: vimage_def image_iff intro!: measure_eq_0I)
apply metis
done
also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
- by (subst setsum.commute) rule
+ by (subst sum.commute) rule
also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
- by (auto simp add: setsum_distrib_left vimage_def intro!: setsum.cong prob_conj_imp1)
+ by (auto simp add: sum_distrib_left vimage_def intro!: sum.cong prob_conj_imp1)
finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
-(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
qed simp_all
@@ -351,8 +351,8 @@
also have "\<dots> =
(\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
unfolding P_def using \<open>K k \<noteq> 0\<close> \<open>k \<in> keys\<close>
- apply (simp add: setsum_cartesian_product' setsum_divide_distrib msgs_def ** cong: conj_cong)
- apply (subst setprod_setsum_distrib_lists[OF M.finite_space]) ..
+ apply (simp add: sum_cartesian_product' sum_divide_distrib msgs_def ** cong: conj_cong)
+ apply (subst setprod_sum_distrib_lists[OF M.finite_space]) ..
finally have "(\<P>(OB ; fst) {(obs, k)}) / K k =
(\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)" . }
note * = this
@@ -373,7 +373,7 @@
have "\<P>(t\<circ>OB ; fst) {(t obs, k)} = (\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)})"
unfolding comp_def
using finite_measure_finite_Union[OF _ _ df]
- by (auto simp add: * intro!: setsum_nonneg)
+ by (auto simp add: * intro!: sum_nonneg)
also have "(\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)}) = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}"
by (simp add: t_eq_imp[OF \<open>k \<in> keys\<close> \<open>K k \<noteq> 0\<close> obs])
finally have "\<P>(t\<circ>OB ; fst) {(t obs, k)} = real (card (?S obs)) * \<P>(OB ; fst) {(obs, k)}" .}
@@ -396,18 +396,18 @@
also have "(\<P>(t\<circ>OB) {t obs}) = (\<Sum>k'\<in>keys. (\<P>(t\<circ>OB|fst) {(t obs, k')}) * (\<P>(fst) {k'}))"
using finite_measure_mono[of "{x. t (OB x) = t obs \<and> fst x = k} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
- apply (simp add: setsum_distribution_cut[of "t\<circ>OB" "t obs" fst])
- apply (auto intro!: setsum.cong simp: subset_eq vimage_def prob_conj_imp1)
+ apply (simp add: sum_distribution_cut[of "t\<circ>OB" "t obs" fst])
+ apply (auto intro!: sum.cong simp: subset_eq vimage_def prob_conj_imp1)
done
also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
using CP_t_K[OF \<open>k\<in>keys\<close> obs] CP_t_K[OF _ obs] \<open>real (card ?S) \<noteq> 0\<close>
- by (simp only: setsum_distrib_left[symmetric] ac_simps
+ by (simp only: sum_distrib_left[symmetric] ac_simps
mult_divide_mult_cancel_left[OF \<open>real (card ?S) \<noteq> 0\<close>]
- cong: setsum.cong)
+ cong: sum.cong)
also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
- using setsum_distribution_cut[of OB obs fst]
- by (auto intro!: setsum.cong simp: prob_conj_imp1 vimage_def)
+ using sum_distribution_cut[of OB obs fst]
+ by (auto intro!: sum.cong simp: prob_conj_imp1 vimage_def)
also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
by (auto simp: vimage_def conj_commute prob_conj_imp2)
finally have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}" . }
@@ -433,7 +433,7 @@
have "\<P>(t\<circ>OB) {t (OB x)} = (\<Sum>obs\<in>?S (OB x). \<P>(OB) {obs})"
unfolding comp_def
using finite_measure_finite_Union[OF _ _ df]
- by (force simp add: * intro!: setsum_nonneg) }
+ by (force simp add: * intro!: sum_nonneg) }
note P_t_sum_P_O = this
txt \<open>Lemma 3\<close>
@@ -441,16 +441,16 @@
unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp
also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t])
- apply (subst setsum.reindex)
+ apply (subst sum.reindex)
apply (fastforce intro!: inj_onI)
apply simp
- apply (subst setsum.Sigma[symmetric, unfolded split_def])
+ apply (subst sum.Sigma[symmetric, unfolded split_def])
using finite_space apply fastforce
using finite_space apply fastforce
- apply (safe intro!: setsum.cong)
+ apply (safe intro!: sum.cong)
using P_t_sum_P_O
- by (simp add: setsum_divide_distrib[symmetric] field_simps **
- setsum_distrib_left[symmetric] setsum_distrib_right[symmetric])
+ by (simp add: sum_divide_distrib[symmetric] field_simps **
+ sum_distrib_left[symmetric] sum_distrib_right[symmetric])
also have "\<dots> = \<H>(fst | t\<circ>OB)"
unfolding conditional_entropy_eq_ce_with_hypothesis
by (simp add: comp_def image_image[symmetric])
--- a/src/HOL/Rat.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Rat.thy Mon Oct 17 14:37:32 2016 +0200
@@ -698,7 +698,7 @@
lemma of_rat_mult: "of_rat (a * b) = of_rat a * of_rat b"
by transfer (simp add: divide_inverse nonzero_inverse_mult_distrib ac_simps)
-lemma of_rat_setsum: "of_rat (\<Sum>a\<in>A. f a) = (\<Sum>a\<in>A. of_rat (f a))"
+lemma of_rat_sum: "of_rat (\<Sum>a\<in>A. f a) = (\<Sum>a\<in>A. of_rat (f a))"
by (induct rule: infinite_finite_induct) (auto simp: of_rat_add)
lemma of_rat_setprod: "of_rat (\<Prod>a\<in>A. f a) = (\<Prod>a\<in>A. of_rat (f a))"
--- a/src/HOL/Real.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Real.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1125,7 +1125,7 @@
subsection \<open>Embedding the Naturals into the Reals\<close>
-lemma real_of_card: "real (card A) = setsum (\<lambda>x. 1) A"
+lemma real_of_card: "real (card A) = sum (\<lambda>x. 1) A"
by simp
lemma nat_less_real_le: "n < m \<longleftrightarrow> real n + 1 \<le> real m"
--- a/src/HOL/Real_Vector_Spaces.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Real_Vector_Spaces.thy Mon Oct 17 14:37:32 2016 +0200
@@ -33,7 +33,7 @@
lemma diff: "f (x - y) = f x - f y"
using add [of x "- y"] by (simp add: minus)
-lemma setsum: "f (setsum g A) = (\<Sum>x\<in>A. f (g x))"
+lemma sum: "f (sum g A) = (\<Sum>x\<in>A. f (g x))"
by (induct A rule: infinite_finite_induct) (simp_all add: zero add)
end
@@ -55,27 +55,27 @@
lemma scale_zero_left [simp]: "scale 0 x = 0"
and scale_minus_left [simp]: "scale (- a) x = - (scale a x)"
and scale_left_diff_distrib [algebra_simps]: "scale (a - b) x = scale a x - scale b x"
- and scale_setsum_left: "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)"
+ and scale_sum_left: "scale (sum f A) x = (\<Sum>a\<in>A. scale (f a) x)"
proof -
interpret s: additive "\<lambda>a. scale a x"
by standard (rule scale_left_distrib)
show "scale 0 x = 0" by (rule s.zero)
show "scale (- a) x = - (scale a x)" by (rule s.minus)
show "scale (a - b) x = scale a x - scale b x" by (rule s.diff)
- show "scale (setsum f A) x = (\<Sum>a\<in>A. scale (f a) x)" by (rule s.setsum)
+ show "scale (sum f A) x = (\<Sum>a\<in>A. scale (f a) x)" by (rule s.sum)
qed
lemma scale_zero_right [simp]: "scale a 0 = 0"
and scale_minus_right [simp]: "scale a (- x) = - (scale a x)"
and scale_right_diff_distrib [algebra_simps]: "scale a (x - y) = scale a x - scale a y"
- and scale_setsum_right: "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))"
+ and scale_sum_right: "scale a (sum f A) = (\<Sum>x\<in>A. scale a (f x))"
proof -
interpret s: additive "\<lambda>x. scale a x"
by standard (rule scale_right_distrib)
show "scale a 0 = 0" by (rule s.zero)
show "scale a (- x) = - (scale a x)" by (rule s.minus)
show "scale a (x - y) = scale a x - scale a y" by (rule s.diff)
- show "scale a (setsum f A) = (\<Sum>x\<in>A. scale a (f x))" by (rule s.setsum)
+ show "scale a (sum f A) = (\<Sum>x\<in>A. scale a (f x))" by (rule s.sum)
qed
lemma scale_eq_0_iff [simp]: "scale a x = 0 \<longleftrightarrow> a = 0 \<or> x = 0"
@@ -154,11 +154,11 @@
lemmas scaleR_zero_left = real_vector.scale_zero_left
lemmas scaleR_minus_left = real_vector.scale_minus_left
lemmas scaleR_diff_left = real_vector.scale_left_diff_distrib
-lemmas scaleR_setsum_left = real_vector.scale_setsum_left
+lemmas scaleR_sum_left = real_vector.scale_sum_left
lemmas scaleR_zero_right = real_vector.scale_zero_right
lemmas scaleR_minus_right = real_vector.scale_minus_right
lemmas scaleR_diff_right = real_vector.scale_right_diff_distrib
-lemmas scaleR_setsum_right = real_vector.scale_setsum_right
+lemmas scaleR_sum_right = real_vector.scale_sum_right
lemmas scaleR_eq_0_iff = real_vector.scale_eq_0_iff
lemmas scaleR_left_imp_eq = real_vector.scale_left_imp_eq
lemmas scaleR_right_imp_eq = real_vector.scale_right_imp_eq
@@ -216,7 +216,7 @@
apply (erule (1) nonzero_inverse_scaleR_distrib)
done
-lemma setsum_constant_scaleR: "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
+lemma sum_constant_scaleR: "(\<Sum>x\<in>A. y) = of_nat (card A) *\<^sub>R y"
for y :: "'a::real_vector"
by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps)
@@ -316,7 +316,7 @@
lemma of_real_mult [simp]: "of_real (x * y) = of_real x * of_real y"
by (simp add: of_real_def mult.commute)
-lemma of_real_setsum[simp]: "of_real (setsum f s) = (\<Sum>x\<in>s. of_real (f x))"
+lemma of_real_sum[simp]: "of_real (sum f s) = (\<Sum>x\<in>s. of_real (f x))"
by (induct s rule: infinite_finite_induct) auto
lemma of_real_setprod[simp]: "of_real (setprod f s) = (\<Prod>x\<in>s. of_real (f x))"
@@ -495,10 +495,10 @@
then show thesis ..
qed
-lemma setsum_in_Reals [intro,simp]: "(\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>) \<Longrightarrow> setsum f s \<in> \<real>"
+lemma sum_in_Reals [intro,simp]: "(\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>) \<Longrightarrow> sum f s \<in> \<real>"
proof (induct s rule: infinite_finite_induct)
case infinite
- then show ?case by (metis Reals_0 setsum.infinite)
+ then show ?case by (metis Reals_0 sum.infinite)
qed simp_all
lemma setprod_in_Reals [intro,simp]: "(\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>) \<Longrightarrow> setprod f s \<in> \<real>"
@@ -848,16 +848,16 @@
shows "norm a \<le> r \<Longrightarrow> norm b \<le> s \<Longrightarrow> norm (a + b) \<le> r + s"
by (metis add_mono_thms_linordered_semiring(1) norm_triangle_ineq order.trans)
-lemma norm_setsum:
+lemma norm_sum:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
- shows "norm (setsum f A) \<le> (\<Sum>i\<in>A. norm (f i))"
+ shows "norm (sum f A) \<le> (\<Sum>i\<in>A. norm (f i))"
by (induct A rule: infinite_finite_induct) (auto intro: norm_triangle_mono)
-lemma setsum_norm_le:
+lemma sum_norm_le:
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
assumes fg: "\<forall>x \<in> S. norm (f x) \<le> g x"
- shows "norm (setsum f S) \<le> setsum g S"
- by (rule order_trans [OF norm_setsum setsum_mono]) (simp add: fg)
+ shows "norm (sum f S) \<le> sum g S"
+ by (rule order_trans [OF norm_sum sum_mono]) (simp add: fg)
lemma abs_norm_cancel [simp]: "\<bar>norm a\<bar> = norm a"
for a :: "'a::real_normed_vector"
@@ -1471,11 +1471,11 @@
lemma diff_right: "prod a (b - b') = prod a b - prod a b'"
by (rule additive.diff [OF additive_right])
-lemma setsum_left: "prod (setsum g S) x = setsum ((\<lambda>i. prod (g i) x)) S"
- by (rule additive.setsum [OF additive_left])
+lemma sum_left: "prod (sum g S) x = sum ((\<lambda>i. prod (g i) x)) S"
+ by (rule additive.sum [OF additive_left])
-lemma setsum_right: "prod x (setsum g S) = setsum ((\<lambda>i. (prod x (g i)))) S"
- by (rule additive.setsum [OF additive_right])
+lemma sum_right: "prod x (sum g S) = sum ((\<lambda>i. (prod x (g i)))) S"
+ by (rule additive.sum [OF additive_right])
lemma bounded_linear_left: "bounded_linear (\<lambda>a. a ** b)"
@@ -1578,7 +1578,7 @@
using bounded_linear_add[of f "\<lambda>x. - g x"] bounded_linear_minus[of g]
by (auto simp add: algebra_simps)
-lemma bounded_linear_setsum:
+lemma bounded_linear_sum:
fixes f :: "'i \<Rightarrow> 'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
shows "(\<And>i. i \<in> I \<Longrightarrow> bounded_linear (f i)) \<Longrightarrow> bounded_linear (\<lambda>x. \<Sum>i\<in>I. f i x)"
by (induct I rule: infinite_finite_induct) (auto intro!: bounded_linear_add)
--- a/src/HOL/Series.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Series.thy Mon Oct 17 14:37:32 2016 +0200
@@ -3,7 +3,7 @@
Copyright : 1998 University of Cambridge
Converted to Isar and polished by lcp
-Converted to setsum and polished yet more by TNN
+Converted to sum and polished yet more by TNN
Additional contributions by Jeremy Avigad
*)
@@ -51,9 +51,9 @@
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_iff_convergent': "summable f \<longleftrightarrow> convergent (\<lambda>n. setsum f {..n})"
+lemma summable_iff_convergent': "summable f \<longleftrightarrow> convergent (\<lambda>n. sum f {..n})"
by (simp_all only: summable_iff_convergent convergent_def
- lessThan_Suc_atMost [symmetric] LIMSEQ_Suc_iff[of "\<lambda>n. setsum f {..<n}"])
+ lessThan_Suc_atMost [symmetric] LIMSEQ_Suc_iff[of "\<lambda>n. sum f {..<n}"])
lemma suminf_eq_lim: "suminf f = lim (\<lambda>n. \<Sum>i<n. f i)"
by (simp add: suminf_def sums_def lim_def)
@@ -64,8 +64,8 @@
lemma summable_zero[simp, intro]: "summable (\<lambda>n. 0)"
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)
+lemma sums_group: "f sums s \<Longrightarrow> 0 < k \<Longrightarrow> (\<lambda>n. sum f {n * k ..< n * k + k}) sums s"
+ apply (simp only: sums_def sum_nat_group tendsto_def eventually_sequentially)
apply safe
apply (erule_tac x=S in allE)
apply safe
@@ -88,22 +88,22 @@
by (auto simp: eventually_at_top_linorder)
define C where "C = (\<Sum>k<N. f k - g k)"
from eventually_ge_at_top[of N]
- have "eventually (\<lambda>n. setsum f {..<n} = C + setsum g {..<n}) sequentially"
+ have "eventually (\<lambda>n. sum f {..<n} = C + sum g {..<n}) sequentially"
proof eventually_elim
case (elim n)
then have "{..<n} = {..<N} \<union> {N..<n}"
by auto
- also have "setsum f ... = setsum f {..<N} + setsum f {N..<n}"
- by (intro setsum.union_disjoint) auto
- also from N have "setsum f {N..<n} = setsum g {N..<n}"
- by (intro setsum.cong) simp_all
- also have "setsum f {..<N} + setsum g {N..<n} = C + (setsum g {..<N} + setsum g {N..<n})"
- unfolding C_def by (simp add: algebra_simps setsum_subtractf)
- also have "setsum g {..<N} + setsum g {N..<n} = setsum g ({..<N} \<union> {N..<n})"
- by (intro setsum.union_disjoint [symmetric]) auto
+ also have "sum f ... = sum f {..<N} + sum f {N..<n}"
+ by (intro sum.union_disjoint) auto
+ also from N have "sum f {N..<n} = sum g {N..<n}"
+ by (intro sum.cong) simp_all
+ also have "sum f {..<N} + sum g {N..<n} = C + (sum g {..<N} + sum g {N..<n})"
+ unfolding C_def by (simp add: algebra_simps sum_subtractf)
+ also have "sum g {..<N} + sum g {N..<n} = sum g ({..<N} \<union> {N..<n})"
+ by (intro sum.union_disjoint [symmetric]) auto
also from elim have "{..<N} \<union> {N..<n} = {..<n}"
by auto
- finally show "setsum f {..<n} = C + setsum g {..<n}" .
+ finally show "sum f {..<n} = C + sum g {..<n}" .
qed
from convergent_cong[OF this] show ?thesis
by (simp add: summable_iff_convergent convergent_add_const_iff)
@@ -114,7 +114,7 @@
and f: "\<And>n. n \<notin> N \<Longrightarrow> f n = 0"
shows "f sums (\<Sum>n\<in>N. f n)"
proof -
- have eq: "setsum f {..<n + Suc (Max N)} = setsum f N" for n
+ have eq: "sum f {..<n + Suc (Max N)} = sum f N" for n
proof (cases "N = {}")
case True
with f have "f = (\<lambda>x. 0)" by auto
@@ -122,7 +122,7 @@
next
case [simp]: False
show ?thesis
- proof (safe intro!: setsum.mono_neutral_right f)
+ proof (safe intro!: sum.mono_neutral_right f)
fix i
assume "i \<in> N"
then have "i \<le> Max N" by simp
@@ -136,7 +136,7 @@
qed
corollary sums_0: "(\<And>n. f n = 0) \<Longrightarrow> (f sums 0)"
- by (metis (no_types) finite.emptyI setsum.empty sums_finite)
+ by (metis (no_types) finite.emptyI sum.empty sums_finite)
lemma summable_finite: "finite N \<Longrightarrow> (\<And>n. n \<notin> N \<Longrightarrow> f n = 0) \<Longrightarrow> summable f"
by (rule sums_summable) (rule sums_finite)
@@ -199,7 +199,7 @@
lemma sums_le: "\<forall>n. f n \<le> g n \<Longrightarrow> f sums s \<Longrightarrow> g sums t \<Longrightarrow> s \<le> t"
for f g :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add,linorder_topology}"
- by (rule LIMSEQ_le) (auto intro: setsum_mono simp: sums_def)
+ by (rule LIMSEQ_le) (auto intro: sum_mono simp: sums_def)
context
fixes f :: "nat \<Rightarrow> 'a::{ordered_comm_monoid_add,linorder_topology}"
@@ -208,13 +208,13 @@
lemma suminf_le: "\<forall>n. f n \<le> g n \<Longrightarrow> summable f \<Longrightarrow> summable g \<Longrightarrow> suminf f \<le> suminf g"
by (auto dest: sums_summable intro: sums_le)
-lemma setsum_le_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> setsum f {..<n} \<le> suminf f"
+lemma sum_le_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> sum f {..<n} \<le> suminf f"
by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto
lemma suminf_nonneg: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 \<le> suminf f"
- using setsum_le_suminf[of 0] by simp
+ using sum_le_suminf[of 0] by simp
-lemma suminf_le_const: "summable f \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
+lemma suminf_le_const: "summable f \<Longrightarrow> (\<And>n. sum f {..<n} \<le> x) \<Longrightarrow> suminf f \<le> x"
by (metis LIMSEQ_le_const2 summable_LIMSEQ)
lemma suminf_eq_zero_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> suminf f = 0 \<longleftrightarrow> (\<forall>n. f n = 0)"
@@ -224,24 +224,24 @@
using summable_LIMSEQ[of f] by simp
then have "\<And>i. (\<Sum>n\<in>{i}. f n) \<le> 0"
proof (rule LIMSEQ_le_const)
- show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> setsum f {..<n}" for i
- using pos by (intro exI[of _ "Suc i"] allI impI setsum_mono2) auto
+ show "\<exists>N. \<forall>n\<ge>N. (\<Sum>n\<in>{i}. f n) \<le> sum f {..<n}" for i
+ using pos by (intro exI[of _ "Suc i"] allI impI sum_mono2) auto
qed
with pos show "\<forall>n. f n = 0"
by (auto intro!: antisym)
qed (metis suminf_zero fun_eq_iff)
lemma suminf_pos_iff: "summable f \<Longrightarrow> \<forall>n. 0 \<le> f n \<Longrightarrow> 0 < suminf f \<longleftrightarrow> (\<exists>i. 0 < f i)"
- using setsum_le_suminf[of 0] suminf_eq_zero_iff by (simp add: less_le)
+ using sum_le_suminf[of 0] suminf_eq_zero_iff by (simp add: less_le)
lemma suminf_pos2:
assumes "summable f" "\<forall>n. 0 \<le> f n" "0 < f i"
shows "0 < suminf f"
proof -
have "0 < (\<Sum>n<Suc i. f n)"
- using assms by (intro setsum_pos2[where i=i]) auto
+ using assms by (intro sum_pos2[where i=i]) auto
also have "\<dots> \<le> suminf f"
- using assms by (intro setsum_le_suminf) auto
+ using assms by (intro sum_le_suminf) auto
finally show ?thesis .
qed
@@ -254,15 +254,15 @@
fixes f :: "nat \<Rightarrow> 'a::{ordered_cancel_comm_monoid_add,linorder_topology}"
begin
-lemma setsum_less_suminf2:
- "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> n \<le> i \<Longrightarrow> 0 < f i \<Longrightarrow> setsum f {..<n} < suminf f"
- using setsum_le_suminf[of f "Suc i"]
- and add_strict_increasing[of "f i" "setsum f {..<n}" "setsum f {..<i}"]
- and setsum_mono2[of "{..<i}" "{..<n}" f]
+lemma sum_less_suminf2:
+ "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 \<le> f m \<Longrightarrow> n \<le> i \<Longrightarrow> 0 < f i \<Longrightarrow> sum f {..<n} < suminf f"
+ using sum_le_suminf[of f "Suc i"]
+ and add_strict_increasing[of "f i" "sum f {..<n}" "sum f {..<i}"]
+ and sum_mono2[of "{..<i}" "{..<n}" f]
by (auto simp: less_imp_le ac_simps)
-lemma setsum_less_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 < f m \<Longrightarrow> setsum f {..<n} < suminf f"
- using setsum_less_suminf2[of n n] by (simp add: less_imp_le)
+lemma sum_less_suminf: "summable f \<Longrightarrow> \<forall>m\<ge>n. 0 < f m \<Longrightarrow> sum f {..<n} < suminf f"
+ using sum_less_suminf2[of n n] by (simp add: less_imp_le)
end
@@ -273,10 +273,10 @@
shows "summable f"
unfolding summable_def sums_def [abs_def]
proof (rule exI LIMSEQ_incseq_SUP)+
- show "bdd_above (range (\<lambda>n. setsum f {..<n}))"
+ show "bdd_above (range (\<lambda>n. sum f {..<n}))"
using le by (auto simp: bdd_above_def)
- show "incseq (\<lambda>n. setsum f {..<n})"
- by (auto simp: mono_def intro!: setsum_mono2)
+ show "incseq (\<lambda>n. sum f {..<n})"
+ by (auto simp: mono_def intro!: sum_mono2)
qed
lemma summableI[intro, simp]: "summable f"
@@ -298,13 +298,13 @@
using assms by (auto intro!: tendsto_add simp: sums_def)
moreover have "(\<Sum>i<n. f (Suc i)) + f 0 = (\<Sum>i<Suc n. f i)" for n
unfolding lessThan_Suc_eq_insert_0
- by (simp add: ac_simps setsum_atLeast1_atMost_eq image_Suc_lessThan)
+ by (simp add: ac_simps sum_atLeast1_atMost_eq image_Suc_lessThan)
ultimately show ?thesis
- by (auto simp: sums_def simp del: setsum_lessThan_Suc intro: LIMSEQ_Suc_iff[THEN iffD1])
+ by (auto simp: sums_def simp del: sum_lessThan_Suc intro: LIMSEQ_Suc_iff[THEN iffD1])
qed
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.distrib tendsto_add)
+ unfolding sums_def by (simp add: sum.distrib 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)
@@ -319,14 +319,14 @@
and I :: "'i set"
begin
-lemma sums_setsum: "(\<And>i. i \<in> I \<Longrightarrow> (f i) sums (x i)) \<Longrightarrow> (\<lambda>n. \<Sum>i\<in>I. f i n) sums (\<Sum>i\<in>I. x i)"
+lemma sums_sum: "(\<And>i. i \<in> I \<Longrightarrow> (f i) sums (x i)) \<Longrightarrow> (\<lambda>n. \<Sum>i\<in>I. f i n) sums (\<Sum>i\<in>I. x i)"
by (induct I rule: infinite_finite_induct) (auto intro!: sums_add)
-lemma suminf_setsum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> (\<Sum>n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. \<Sum>n. f i n)"
- using sums_unique[OF sums_setsum, OF summable_sums] by simp
+lemma suminf_sum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> (\<Sum>n. \<Sum>i\<in>I. f i n) = (\<Sum>i\<in>I. \<Sum>n. f i n)"
+ using sums_unique[OF sums_sum, OF summable_sums] by simp
-lemma summable_setsum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> summable (\<lambda>n. \<Sum>i\<in>I. f i n)"
- using sums_summable[OF sums_setsum[OF summable_sums]] .
+lemma summable_sum: "(\<And>i. i \<in> I \<Longrightarrow> summable (f i)) \<Longrightarrow> summable (\<lambda>n. \<Sum>i\<in>I. f i n)"
+ using sums_summable[OF sums_sum[OF summable_sums]] .
end
@@ -341,7 +341,7 @@
have "f sums (s + f 0) \<longleftrightarrow> (\<lambda>i. \<Sum>j<Suc i. f j) \<longlonglongrightarrow> 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) \<longlonglongrightarrow> s + f 0"
- by (simp add: ac_simps lessThan_Suc_eq_insert_0 image_Suc_lessThan setsum_atLeast1_atMost_eq)
+ by (simp add: ac_simps lessThan_Suc_eq_insert_0 image_Suc_lessThan sum_atLeast1_atMost_eq)
also have "\<dots> \<longleftrightarrow> (\<lambda>n. f (Suc n)) sums s"
proof
assume "(\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) \<longlonglongrightarrow> s + f 0"
@@ -372,7 +372,7 @@
begin
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)
+ unfolding sums_def by (simp add: sum_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)
@@ -381,7 +381,7 @@
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)
+ unfolding sums_def by (simp add: sum_negf tendsto_minus)
lemma summable_minus: "summable f \<Longrightarrow> summable (\<lambda>n. - f n)"
unfolding summable_def by (auto intro: sums_minus)
@@ -433,7 +433,7 @@
shows "\<exists>N. \<forall>n\<ge>N. norm (\<Sum>i. f (i + n)) < r"
proof -
from LIMSEQ_D[OF summable_LIMSEQ[OF \<open>summable f\<close>] \<open>0 < r\<close>]
- obtain N :: nat where "\<forall> n \<ge> N. norm (setsum f {..<n} - suminf f) < r"
+ obtain N :: nat where "\<forall> n \<ge> N. norm (sum f {..<n} - suminf f) < r"
by auto
then show ?thesis
by (auto simp: norm_minus_commute suminf_minus_initial_segment[OF \<open>summable f\<close>])
@@ -467,7 +467,7 @@
by (auto dest: summable_minus) (* used two ways, hence must be outside the context above *)
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)
+ unfolding sums_def by (drule tendsto) (simp only: sum)
lemma (in bounded_linear) summable: "summable (\<lambda>n. X n) \<Longrightarrow> summable (\<lambda>n. f (X n))"
unfolding summable_def by (auto intro: sums)
@@ -497,7 +497,7 @@
(auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_real_sequentially)
then have "\<not> convergent (\<lambda>n. norm (\<Sum>k<n. c))"
by (intro filterlim_at_infinity_imp_not_convergent filterlim_at_top_imp_at_infinity)
- (simp_all add: setsum_constant_scaleR)
+ (simp_all add: sum_constant_scaleR)
then show ?thesis
unfolding summable_iff_convergent using convergent_norm by blast
qed
@@ -546,7 +546,7 @@
lemma sums_of_real_iff:
"(\<lambda>n. of_real (f n) :: 'a::real_normed_div_algebra) sums of_real c \<longleftrightarrow> f sums c"
- by (simp add: sums_def of_real_setsum[symmetric] tendsto_of_real_iff del: of_real_setsum)
+ by (simp add: sums_def of_real_sum[symmetric] tendsto_of_real_iff del: of_real_sum)
subsection \<open>Infinite summability on real normed fields\<close>
@@ -628,7 +628,7 @@
unfolding sums_def
proof (subst LIMSEQ_Suc_iff [symmetric])
have "(\<lambda>n. \<Sum>k<Suc n. f (Suc k) - f k) = (\<lambda>n. f (Suc n) - f 0)"
- by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] setsum_Suc_diff)
+ by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] sum_Suc_diff)
also have "\<dots> \<longlonglongrightarrow> c - f 0"
by (intro tendsto_diff LIMSEQ_Suc[OF assms] tendsto_const)
finally show "(\<lambda>n. \<Sum>n<Suc n. f (Suc n) - f n) \<longlonglongrightarrow> c - f 0" .
@@ -657,7 +657,7 @@
text \<open>Cauchy-type criterion for convergence of series (c.f. Harrison).\<close>
-lemma summable_Cauchy: "summable f \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n. norm (setsum f {m..<n}) < e)"
+lemma summable_Cauchy: "summable f \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m\<ge>N. \<forall>n. norm (sum f {m..<n}) < e)"
for f :: "nat \<Rightarrow> 'a::banach"
apply (simp only: summable_iff_convergent Cauchy_convergent_iff [symmetric] Cauchy_iff)
apply safe
@@ -672,7 +672,7 @@
apply (drule (1) mp)
apply (drule_tac x="m" in spec)
apply (drule (1) mp)
- apply (simp_all add: setsum_diff [symmetric])
+ apply (simp_all add: sum_diff [symmetric])
apply (drule spec)
apply (drule (1) mp)
apply (erule exE)
@@ -680,7 +680,7 @@
apply 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])
+ apply (simp_all add: sum_diff [symmetric])
done
context
@@ -698,13 +698,13 @@
apply safe
apply (drule_tac x="m" in spec)
apply safe
- apply (rule order_le_less_trans [OF norm_setsum])
+ apply (rule order_le_less_trans [OF norm_sum])
apply (rule order_le_less_trans [OF abs_ge_self])
apply simp
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)
+ by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_sum)
text \<open>Comparison tests.\<close>
@@ -721,9 +721,9 @@
apply (rotate_tac 2)
apply (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)
+ apply (rule norm_sum)
+ apply (rule_tac y = "sum g {m..<n}" in order_le_less_trans)
+ apply (auto intro: sum_mono simp add: abs_less_iff)
done
lemma summable_comparison_test_ev:
@@ -841,58 +841,58 @@
let ?g = "\<lambda>(i,j). a i * b j"
let ?f = "\<lambda>(i,j). norm (a i) * norm (b j)"
have f_nonneg: "\<And>x. 0 \<le> ?f x" by auto
- then have norm_setsum_f: "\<And>A. norm (setsum ?f A) = setsum ?f A"
+ then have norm_sum_f: "\<And>A. norm (sum ?f A) = sum ?f A"
unfolding real_norm_def
- by (simp only: abs_of_nonneg setsum_nonneg [rule_format])
+ by (simp only: abs_of_nonneg sum_nonneg [rule_format])
have "(\<lambda>n. (\<Sum>k<n. a k) * (\<Sum>k<n. b k)) \<longlonglongrightarrow> (\<Sum>k. a k) * (\<Sum>k. b k)"
by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b])
- then have 1: "(\<lambda>n. setsum ?g (?S1 n)) \<longlonglongrightarrow> (\<Sum>k. a k) * (\<Sum>k. b k)"
- by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)
+ then have 1: "(\<lambda>n. sum ?g (?S1 n)) \<longlonglongrightarrow> (\<Sum>k. a k) * (\<Sum>k. b k)"
+ by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan)
have "(\<lambda>n. (\<Sum>k<n. norm (a k)) * (\<Sum>k<n. norm (b k))) \<longlonglongrightarrow> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
using a b by (intro tendsto_mult summable_LIMSEQ)
- then have "(\<lambda>n. setsum ?f (?S1 n)) \<longlonglongrightarrow> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
- by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)
- then have "convergent (\<lambda>n. setsum ?f (?S1 n))"
+ then have "(\<lambda>n. sum ?f (?S1 n)) \<longlonglongrightarrow> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
+ by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan)
+ then have "convergent (\<lambda>n. sum ?f (?S1 n))"
by (rule convergentI)
- then have Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"
+ then have Cauchy: "Cauchy (\<lambda>n. sum ?f (?S1 n))"
by (rule convergent_Cauchy)
- have "Zfun (\<lambda>n. setsum ?f (?S1 n - ?S2 n)) sequentially"
- proof (rule ZfunI, simp only: eventually_sequentially norm_setsum_f)
+ have "Zfun (\<lambda>n. sum ?f (?S1 n - ?S2 n)) sequentially"
+ proof (rule ZfunI, simp only: eventually_sequentially norm_sum_f)
fix r :: real
assume r: "0 < r"
from CauchyD [OF Cauchy r] obtain N
- where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (setsum ?f (?S1 m) - setsum ?f (?S1 n)) < r" ..
- then have "\<And>m n. N \<le> n \<Longrightarrow> n \<le> m \<Longrightarrow> norm (setsum ?f (?S1 m - ?S1 n)) < r"
- by (simp only: setsum_diff finite_S1 S1_mono)
- then have N: "\<And>m n. N \<le> n \<Longrightarrow> n \<le> m \<Longrightarrow> setsum ?f (?S1 m - ?S1 n) < r"
- by (simp only: norm_setsum_f)
- show "\<exists>N. \<forall>n\<ge>N. setsum ?f (?S1 n - ?S2 n) < r"
+ where "\<forall>m\<ge>N. \<forall>n\<ge>N. norm (sum ?f (?S1 m) - sum ?f (?S1 n)) < r" ..
+ then have "\<And>m n. N \<le> n \<Longrightarrow> n \<le> m \<Longrightarrow> norm (sum ?f (?S1 m - ?S1 n)) < r"
+ by (simp only: sum_diff finite_S1 S1_mono)
+ then have N: "\<And>m n. N \<le> n \<Longrightarrow> n \<le> m \<Longrightarrow> sum ?f (?S1 m - ?S1 n) < r"
+ by (simp only: norm_sum_f)
+ show "\<exists>N. \<forall>n\<ge>N. sum ?f (?S1 n - ?S2 n) < r"
proof (intro exI allI impI)
fix n
assume "2 * N \<le> n"
then have n: "N \<le> n div 2" by simp
- have "setsum ?f (?S1 n - ?S2 n) \<le> setsum ?f (?S1 n - ?S1 (n div 2))"
- by (intro setsum_mono2 finite_Diff finite_S1 f_nonneg Diff_mono subset_refl S1_le_S2)
+ have "sum ?f (?S1 n - ?S2 n) \<le> sum ?f (?S1 n - ?S1 (n div 2))"
+ by (intro sum_mono2 finite_Diff finite_S1 f_nonneg Diff_mono subset_refl S1_le_S2)
also have "\<dots> < r"
using n div_le_dividend by (rule N)
- finally show "setsum ?f (?S1 n - ?S2 n) < r" .
+ finally show "sum ?f (?S1 n - ?S2 n) < r" .
qed
qed
- then have "Zfun (\<lambda>n. setsum ?g (?S1 n - ?S2 n)) sequentially"
+ then have "Zfun (\<lambda>n. sum ?g (?S1 n - ?S2 n)) sequentially"
apply (rule Zfun_le [rule_format])
- apply (simp only: norm_setsum_f)
- apply (rule order_trans [OF norm_setsum setsum_mono])
+ apply (simp only: norm_sum_f)
+ apply (rule order_trans [OF norm_sum sum_mono])
apply (auto simp add: norm_mult_ineq)
done
- then have 2: "(\<lambda>n. setsum ?g (?S1 n) - setsum ?g (?S2 n)) \<longlonglongrightarrow> 0"
+ then have 2: "(\<lambda>n. sum ?g (?S1 n) - sum ?g (?S2 n)) \<longlonglongrightarrow> 0"
unfolding tendsto_Zfun_iff diff_0_right
- by (simp only: setsum_diff finite_S1 S2_le_S1)
- with 1 have "(\<lambda>n. setsum ?g (?S2 n)) \<longlonglongrightarrow> (\<Sum>k. a k) * (\<Sum>k. b k)"
+ by (simp only: sum_diff finite_S1 S2_le_S1)
+ with 1 have "(\<lambda>n. sum ?g (?S2 n)) \<longlonglongrightarrow> (\<Sum>k. a k) * (\<Sum>k. b k)"
by (rule Lim_transform2)
then show ?thesis
- by (simp only: sums_def setsum_triangle_reindex)
+ by (simp only: sums_def sum_triangle_reindex)
qed
lemma Cauchy_product:
@@ -1021,7 +1021,7 @@
with m have "norm ((\<Sum>k<Suc n. f k) - (\<Sum>k<m. f k)) < e"
by (intro N) simp_all
also from True have "(\<Sum>k<Suc n. f k) - (\<Sum>k<m. f k) = (\<Sum>k=m..n. f k)"
- by (subst setsum_diff [symmetric]) (simp_all add: setsum_last_plus)
+ by (subst sum_diff [symmetric]) (simp_all add: sum_last_plus)
finally show ?thesis .
next
case False
@@ -1059,12 +1059,12 @@
then obtain m where n: "\<And>n'. n' < n \<Longrightarrow> g n' < m"
by blast
- have "(\<Sum>i<n. f (g i)) = setsum f (g ` {..<n})"
- by (simp add: setsum.reindex)
+ have "(\<Sum>i<n. f (g i)) = sum f (g ` {..<n})"
+ by (simp add: sum.reindex)
also have "\<dots> \<le> (\<Sum>i<m. f i)"
- by (rule setsum_mono3) (auto simp add: pos n[rule_format])
+ by (rule sum_mono3) (auto simp add: pos n[rule_format])
also have "\<dots> \<le> suminf f"
- using \<open>summable f\<close> by (rule setsum_le_suminf) (simp add: pos)
+ using \<open>summable f\<close> by (rule sum_le_suminf) (simp add: pos)
finally show "(\<Sum>i<n. (f \<circ> g) i) \<le> suminf f"
by simp
qed
@@ -1093,14 +1093,14 @@
then obtain m where n: "\<And>n'. g n' < n \<Longrightarrow> n' < m"
by blast
have "(\<Sum>i<n. f i) = (\<Sum>i\<in>{..<n} \<inter> range g. f i)"
- using f by(auto intro: setsum.mono_neutral_cong_right)
+ using f by(auto intro: sum.mono_neutral_cong_right)
also have "\<dots> = (\<Sum>i\<in>g -` {..<n}. (f \<circ> g) i)"
- by (rule setsum.reindex_cong[where l=g])(auto)
+ by (rule sum.reindex_cong[where l=g])(auto)
also have "\<dots> \<le> (\<Sum>i<m. (f \<circ> g) i)"
- by (rule setsum_mono3)(auto simp add: pos n)
+ by (rule sum_mono3)(auto simp add: pos n)
also have "\<dots> \<le> suminf (f \<circ> g)"
- using \<open>summable (f \<circ> g)\<close> by (rule setsum_le_suminf) (simp add: pos)
- finally show "setsum f {..<n} \<le> suminf (f \<circ> g)" .
+ using \<open>summable (f \<circ> g)\<close> by (rule sum_le_suminf) (simp add: pos)
+ finally show "sum f {..<n} \<le> suminf (f \<circ> g)" .
qed
with le show "suminf (f \<circ> g) = suminf f"
by (rule antisym)
@@ -1117,9 +1117,9 @@
proof
fix n :: nat
from subseq have "(\<Sum>k<n. f (g k)) = (\<Sum>k\<in>g`{..<n}. f k)"
- by (subst setsum.reindex) (auto intro: subseq_imp_inj_on)
+ by (subst sum.reindex) (auto intro: subseq_imp_inj_on)
also from subseq have "\<dots> = (\<Sum>k<g n. f k)"
- by (intro setsum.mono_neutral_left ballI zero)
+ by (intro sum.mono_neutral_left ballI zero)
(auto dest: subseq_strict_mono simp: strict_mono_less strict_mono_less_eq)
finally show "(\<Sum>k<n. f (g k)) = (\<Sum>k<g n. f k)" .
qed
@@ -1158,9 +1158,9 @@
by (rule zero)
}
with g_inv_least' g_inv have "(\<Sum>k<n. f k) = (\<Sum>k\<in>g`{..<g_inv n}. f k)"
- by (intro setsum.mono_neutral_right) auto
+ by (intro sum.mono_neutral_right) auto
also from subseq have "\<dots> = (\<Sum>k<g_inv n. f (g k))"
- using subseq_imp_inj_on by (subst setsum.reindex) simp_all
+ using subseq_imp_inj_on by (subst sum.reindex) simp_all
finally show "(\<Sum>k<n. f k) = (\<Sum>k<g_inv n. f (g k))" .
qed
also {
--- a/src/HOL/Set_Interval.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Set_Interval.thy Mon Oct 17 14:37:32 2016 +0200
@@ -1408,7 +1408,7 @@
qed
-subsection \<open>Lemmas useful with the summation operator setsum\<close>
+subsection \<open>Lemmas useful with the summation operator sum\<close>
text \<open>For examples, see Algebra/poly/UnivPoly2.thy\<close>
@@ -1614,32 +1614,32 @@
subsection \<open>Summation indexed over intervals\<close>
syntax (ASCII)
- "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
- "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10)
- "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<_./ _)" [0,0,10] 10)
- "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<=_./ _)" [0,0,10] 10)
+ "_from_to_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _.._./ _)" [0,0,0,10] 10)
+ "_from_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _ = _..<_./ _)" [0,0,0,10] 10)
+ "_upt_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<_./ _)" [0,0,10] 10)
+ "_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(SUM _<=_./ _)" [0,0,10] 10)
syntax (latex_sum output)
- "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ "_from_to_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
("(3\<^latex>\<open>$\\sum_{\<close>_ = _\<^latex>\<open>}^{\<close>_\<^latex>\<open>}$\<close> _)" [0,0,0,10] 10)
- "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ "_from_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
("(3\<^latex>\<open>$\\sum_{\<close>_ = _\<^latex>\<open>}^{<\<close>_\<^latex>\<open>}$\<close> _)" [0,0,0,10] 10)
- "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ "_upt_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
("(3\<^latex>\<open>$\\sum_{\<close>_ < _\<^latex>\<open>}$\<close> _)" [0,0,10] 10)
- "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
+ "_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b"
("(3\<^latex>\<open>$\\sum_{\<close>_ \<le> _\<^latex>\<open>}$\<close> _)" [0,0,10] 10)
syntax
- "_from_to_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
- "_from_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
- "_upt_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
- "_upto_setsum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
+ "_from_to_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _.._./ _)" [0,0,0,10] 10)
+ "_from_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_ = _..<_./ _)" [0,0,0,10] 10)
+ "_upt_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_<_./ _)" [0,0,10] 10)
+ "_upto_sum" :: "idt \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b" ("(3\<Sum>_\<le>_./ _)" [0,0,10] 10)
translations
- "\<Sum>x=a..b. t" == "CONST setsum (\<lambda>x. t) {a..b}"
- "\<Sum>x=a..<b. t" == "CONST setsum (\<lambda>x. t) {a..<b}"
- "\<Sum>i\<le>n. t" == "CONST setsum (\<lambda>i. t) {..n}"
- "\<Sum>i<n. t" == "CONST setsum (\<lambda>i. t) {..<n}"
+ "\<Sum>x=a..b. t" == "CONST sum (\<lambda>x. t) {a..b}"
+ "\<Sum>x=a..<b. t" == "CONST sum (\<lambda>x. t) {a..<b}"
+ "\<Sum>i\<le>n. t" == "CONST sum (\<lambda>i. t) {..n}"
+ "\<Sum>i<n. t" == "CONST sum (\<lambda>i. t) {..<n}"
text\<open>The above introduces some pretty alternative syntaxes for
summation over intervals:
@@ -1661,42 +1661,42 @@
works well with italic-style formulae, not tt-style.
Note that for uniformity on @{typ nat} it is better to use
-@{term"\<Sum>x::nat=0..<n. e"} rather than \<open>\<Sum>x<n. e\<close>: \<open>setsum\<close> may
+@{term"\<Sum>x::nat=0..<n. e"} rather than \<open>\<Sum>x<n. e\<close>: \<open>sum\<close> may
not provide all lemmas available for @{term"{m..<n}"} also in the
special form for @{term"{..<n}"}.\<close>
text\<open>This congruence rule should be used for sums over intervals as
-the standard theorem @{text[source]setsum.cong} does not work well
+the standard theorem @{text[source]sum.cong} does not work well
with the simplifier who adds the unsimplified premise @{term"x:B"} to
the context.\<close>
-lemmas setsum_ivl_cong = setsum.ivl_cong
+lemmas sum_ivl_cong = sum.ivl_cong
(* FIXME why are the following simp rules but the corresponding eqns
on intervals are not? *)
-lemma setsum_atMost_Suc [simp]:
+lemma sum_atMost_Suc [simp]:
"(\<Sum>i \<le> Suc n. f i) = (\<Sum>i \<le> n. f i) + f (Suc n)"
by (simp add: atMost_Suc ac_simps)
-lemma setsum_lessThan_Suc [simp]:
+lemma sum_lessThan_Suc [simp]:
"(\<Sum>i < Suc n. f i) = (\<Sum>i < n. f i) + f n"
by (simp add: lessThan_Suc ac_simps)
-lemma setsum_cl_ivl_Suc [simp]:
- "setsum f {m..Suc n} = (if Suc n < m then 0 else setsum f {m..n} + f(Suc n))"
+lemma sum_cl_ivl_Suc [simp]:
+ "sum f {m..Suc n} = (if Suc n < m then 0 else sum f {m..n} + f(Suc n))"
by (auto simp: ac_simps atLeastAtMostSuc_conv)
-lemma setsum_op_ivl_Suc [simp]:
- "setsum f {m..<Suc n} = (if n < m then 0 else setsum f {m..<n} + f(n))"
+lemma sum_op_ivl_Suc [simp]:
+ "sum f {m..<Suc n} = (if n < m then 0 else sum f {m..<n} + f(n))"
by (auto simp: ac_simps atLeastLessThanSuc)
(*
-lemma setsum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==>
+lemma sum_cl_ivl_add_one_nat: "(n::nat) <= m + 1 ==>
(\<Sum>i=n..m+1. f i) = (\<Sum>i=n..m. f i) + f(m + 1)"
by (auto simp:ac_simps atLeastAtMostSuc_conv)
*)
-lemma setsum_head:
+lemma sum_head:
fixes n :: nat
assumes mn: "m <= n"
shows "(\<Sum>x\<in>{m..n}. P x) = P m + (\<Sum>x\<in>{m<..n}. P x)" (is "?lhs = ?rhs")
@@ -1710,93 +1710,93 @@
finally show ?thesis .
qed
-lemma setsum_head_Suc:
- "m \<le> n \<Longrightarrow> setsum f {m..n} = f m + setsum f {Suc m..n}"
-by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost)
-
-lemma setsum_head_upt_Suc:
- "m < n \<Longrightarrow> setsum f {m..<n} = f m + setsum f {Suc m..<n}"
-apply(insert setsum_head_Suc[of m "n - Suc 0" f])
+lemma sum_head_Suc:
+ "m \<le> n \<Longrightarrow> sum f {m..n} = f m + sum f {Suc m..n}"
+by (simp add: sum_head atLeastSucAtMost_greaterThanAtMost)
+
+lemma sum_head_upt_Suc:
+ "m < n \<Longrightarrow> sum f {m..<n} = f m + sum f {Suc m..<n}"
+apply(insert sum_head_Suc[of m "n - Suc 0" f])
apply (simp add: atLeastLessThanSuc_atLeastAtMost[symmetric] algebra_simps)
done
-lemma setsum_ub_add_nat: assumes "(m::nat) \<le> n + 1"
- shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}"
+lemma sum_ub_add_nat: assumes "(m::nat) \<le> n + 1"
+ shows "sum f {m..n + p} = sum f {m..n} + sum f {n + 1..n + p}"
proof-
have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using \<open>m \<le> n+1\<close> by auto
- thus ?thesis by (auto simp: ivl_disj_int setsum.union_disjoint
+ thus ?thesis by (auto simp: ivl_disj_int sum.union_disjoint
atLeastSucAtMost_greaterThanAtMost)
qed
-lemmas setsum_add_nat_ivl = setsum.atLeast_lessThan_concat
-
-lemma setsum_diff_nat_ivl:
+lemmas sum_add_nat_ivl = sum.atLeast_lessThan_concat
+
+lemma sum_diff_nat_ivl:
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
shows "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
- setsum f {m..<p} - setsum f {m..<n} = setsum f {n..<p}"
-using setsum_add_nat_ivl [of m n p f,symmetric]
+ sum f {m..<p} - sum f {m..<n} = sum f {n..<p}"
+using sum_add_nat_ivl [of m n p f,symmetric]
apply (simp add: ac_simps)
done
-lemma setsum_natinterval_difff:
+lemma sum_natinterval_difff:
fixes f:: "nat \<Rightarrow> ('a::ab_group_add)"
- shows "setsum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
+ shows "sum (\<lambda>k. f k - f(k + 1)) {(m::nat) .. n} =
(if m <= n then f m - f(n + 1) else 0)"
by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
-lemma setsum_nat_group: "(\<Sum>m<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {..< n * k}"
+lemma sum_nat_group: "(\<Sum>m<n::nat. sum f {m * k ..< m*k + k}) = sum 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])
+ apply (simp_all add: sum_add_nat_ivl add.commute atLeast0LessThan[symmetric])
done
-lemma setsum_triangle_reindex:
+lemma sum_triangle_reindex:
fixes n :: nat
shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k<n. \<Sum>i\<le>k. f i (k - i))"
- apply (simp add: setsum.Sigma)
- apply (rule setsum.reindex_bij_witness[where j="\<lambda>(i, j). (i+j, i)" and i="\<lambda>(k, i). (i, k - i)"])
+ apply (simp add: sum.Sigma)
+ apply (rule sum.reindex_bij_witness[where j="\<lambda>(i, j). (i+j, i)" and i="\<lambda>(k, i). (i, k - i)"])
apply auto
done
-lemma setsum_triangle_reindex_eq:
+lemma sum_triangle_reindex_eq:
fixes n :: nat
shows "(\<Sum>(i,j)\<in>{(i,j). i+j \<le> n}. f i j) = (\<Sum>k\<le>n. \<Sum>i\<le>k. f i (k - i))"
-using setsum_triangle_reindex [of f "Suc n"]
+using sum_triangle_reindex [of f "Suc n"]
by (simp only: Nat.less_Suc_eq_le lessThan_Suc_atMost)
-lemma nat_diff_setsum_reindex: "(\<Sum>i<n. f (n - Suc i)) = (\<Sum>i<n. f i)"
- by (rule setsum.reindex_bij_witness[where i="\<lambda>i. n - Suc i" and j="\<lambda>i. n - Suc i"]) auto
+lemma nat_diff_sum_reindex: "(\<Sum>i<n. f (n - Suc i)) = (\<Sum>i<n. f i)"
+ by (rule sum.reindex_bij_witness[where i="\<lambda>i. n - Suc i" and j="\<lambda>i. n - Suc i"]) auto
subsubsection \<open>Shifting bounds\<close>
-lemma setsum_shift_bounds_nat_ivl:
- "setsum f {m+k..<n+k} = setsum (%i. f(i + k)){m..<n::nat}"
+lemma sum_shift_bounds_nat_ivl:
+ "sum f {m+k..<n+k} = sum (%i. f(i + k)){m..<n::nat}"
by (induct "n", auto simp:atLeastLessThanSuc)
-lemma setsum_shift_bounds_cl_nat_ivl:
- "setsum f {m+k..n+k} = setsum (%i. f(i + k)){m..n::nat}"
- by (rule setsum.reindex_bij_witness[where i="\<lambda>i. i + k" and j="\<lambda>i. i - k"]) auto
-
-corollary setsum_shift_bounds_cl_Suc_ivl:
- "setsum f {Suc m..Suc n} = setsum (%i. f(Suc i)){m..n}"
-by (simp add:setsum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
-
-corollary setsum_shift_bounds_Suc_ivl:
- "setsum f {Suc m..<Suc n} = setsum (%i. f(Suc i)){m..<n}"
-by (simp add:setsum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
-
-lemma setsum_shift_lb_Suc0_0:
- "f(0::nat) = (0::nat) \<Longrightarrow> setsum f {Suc 0..k} = setsum f {0..k}"
-by(simp add:setsum_head_Suc)
-
-lemma setsum_shift_lb_Suc0_0_upt:
- "f(0::nat) = 0 \<Longrightarrow> setsum f {Suc 0..<k} = setsum f {0..<k}"
+lemma sum_shift_bounds_cl_nat_ivl:
+ "sum f {m+k..n+k} = sum (%i. f(i + k)){m..n::nat}"
+ by (rule sum.reindex_bij_witness[where i="\<lambda>i. i + k" and j="\<lambda>i. i - k"]) auto
+
+corollary sum_shift_bounds_cl_Suc_ivl:
+ "sum f {Suc m..Suc n} = sum (%i. f(Suc i)){m..n}"
+by (simp add:sum_shift_bounds_cl_nat_ivl[where k="Suc 0", simplified])
+
+corollary sum_shift_bounds_Suc_ivl:
+ "sum f {Suc m..<Suc n} = sum (%i. f(Suc i)){m..<n}"
+by (simp add:sum_shift_bounds_nat_ivl[where k="Suc 0", simplified])
+
+lemma sum_shift_lb_Suc0_0:
+ "f(0::nat) = (0::nat) \<Longrightarrow> sum f {Suc 0..k} = sum f {0..k}"
+by(simp add:sum_head_Suc)
+
+lemma sum_shift_lb_Suc0_0_upt:
+ "f(0::nat) = 0 \<Longrightarrow> sum f {Suc 0..<k} = sum f {0..<k}"
apply(cases k)apply simp
-apply(simp add:setsum_head_upt_Suc)
+apply(simp add:sum_head_upt_Suc)
done
-lemma setsum_atMost_Suc_shift:
+lemma sum_atMost_Suc_shift:
fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add"
shows "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
proof (induct n)
@@ -1804,71 +1804,71 @@
next
case (Suc n) note IH = this
have "(\<Sum>i\<le>Suc (Suc n). f i) = (\<Sum>i\<le>Suc n. f i) + f (Suc (Suc n))"
- by (rule setsum_atMost_Suc)
+ by (rule sum_atMost_Suc)
also have "(\<Sum>i\<le>Suc n. f i) = f 0 + (\<Sum>i\<le>n. f (Suc i))"
by (rule IH)
also have "f 0 + (\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) =
f 0 + ((\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)))"
by (rule add.assoc)
also have "(\<Sum>i\<le>n. f (Suc i)) + f (Suc (Suc n)) = (\<Sum>i\<le>Suc n. f (Suc i))"
- by (rule setsum_atMost_Suc [symmetric])
+ by (rule sum_atMost_Suc [symmetric])
finally show ?case .
qed
-lemma setsum_lessThan_Suc_shift:
+lemma sum_lessThan_Suc_shift:
"(\<Sum>i<Suc n. f i) = f 0 + (\<Sum>i<n. f (Suc i))"
by (induction n) (simp_all add: add_ac)
-lemma setsum_atMost_shift:
+lemma sum_atMost_shift:
fixes f :: "nat \<Rightarrow> 'a::comm_monoid_add"
shows "(\<Sum>i\<le>n. f i) = f 0 + (\<Sum>i<n. f (Suc i))"
-by (metis atLeast0AtMost atLeast0LessThan atLeastLessThanSuc_atLeastAtMost atLeastSucAtMost_greaterThanAtMost le0 setsum_head setsum_shift_bounds_Suc_ivl)
-
-lemma setsum_last_plus: fixes n::nat shows "m <= n \<Longrightarrow> (\<Sum>i = m..n. f i) = f n + (\<Sum>i = m..<n. f i)"
+by (metis atLeast0AtMost atLeast0LessThan atLeastLessThanSuc_atLeastAtMost atLeastSucAtMost_greaterThanAtMost le0 sum_head sum_shift_bounds_Suc_ivl)
+
+lemma sum_last_plus: fixes n::nat shows "m <= n \<Longrightarrow> (\<Sum>i = m..n. f i) = f n + (\<Sum>i = m..<n. f i)"
by (cases n) (auto simp: atLeastLessThanSuc_atLeastAtMost add.commute)
-lemma setsum_Suc_diff:
+lemma sum_Suc_diff:
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
assumes "m \<le> Suc n"
shows "(\<Sum>i = m..n. f(Suc i) - f i) = f (Suc n) - f m"
using assms by (induct n) (auto simp: le_Suc_eq)
-lemma nested_setsum_swap:
+lemma nested_sum_swap:
"(\<Sum>i = 0..n. (\<Sum>j = 0..<i. a i j)) = (\<Sum>j = 0..<n. \<Sum>i = Suc j..n. a i j)"
- by (induction n) (auto simp: setsum.distrib)
-
-lemma nested_setsum_swap':
+ by (induction n) (auto simp: sum.distrib)
+
+lemma nested_sum_swap':
"(\<Sum>i\<le>n. (\<Sum>j<i. a i j)) = (\<Sum>j<n. \<Sum>i = Suc j..n. a i j)"
- by (induction n) (auto simp: setsum.distrib)
-
-lemma setsum_atLeast1_atMost_eq:
- "setsum f {Suc 0..n} = (\<Sum>k<n. f (Suc k))"
+ by (induction n) (auto simp: sum.distrib)
+
+lemma sum_atLeast1_atMost_eq:
+ "sum f {Suc 0..n} = (\<Sum>k<n. f (Suc k))"
proof -
- have "setsum f {Suc 0..n} = setsum f (Suc ` {..<n})"
+ have "sum f {Suc 0..n} = sum f (Suc ` {..<n})"
by (simp add: image_Suc_lessThan)
also have "\<dots> = (\<Sum>k<n. f (Suc k))"
- by (simp add: setsum.reindex)
+ by (simp add: sum.reindex)
finally show ?thesis .
qed
subsubsection \<open>Telescoping\<close>
-lemma setsum_telescope:
+lemma sum_telescope:
fixes f::"nat \<Rightarrow> 'a::ab_group_add"
- shows "setsum (\<lambda>i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)"
+ shows "sum (\<lambda>i. f i - f (Suc i)) {.. i} = f 0 - f (Suc i)"
by (induct i) simp_all
-lemma setsum_telescope'':
+lemma sum_telescope'':
assumes "m \<le> n"
shows "(\<Sum>k\<in>{Suc m..n}. f k - f (k - 1)) = f n - (f m :: 'a :: ab_group_add)"
by (rule dec_induct[OF assms]) (simp_all add: algebra_simps)
-lemma setsum_lessThan_telescope:
+lemma sum_lessThan_telescope:
"(\<Sum>n<m. f (Suc n) - f n :: 'a :: ab_group_add) = f m - f 0"
by (induction m) (simp_all add: algebra_simps)
-lemma setsum_lessThan_telescope':
+lemma sum_lessThan_telescope':
"(\<Sum>n<m. f n - f (Suc n) :: 'a :: ab_group_add) = f 0 - f m"
by (induction m) (simp_all add: algebra_simps)
@@ -1885,7 +1885,7 @@
ultimately show ?thesis by simp
qed
-lemma diff_power_eq_setsum:
+lemma diff_power_eq_sum:
fixes y :: "'a::{comm_ring,monoid_mult}"
shows
"x ^ (Suc n) - y ^ (Suc n) =
@@ -1901,32 +1901,32 @@
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<Suc (Suc n). x ^ p * y ^ (Suc n - p))"
- by (simp add: field_simps Suc_diff_le setsum_distrib_right setsum_distrib_left)
+ by (simp add: field_simps Suc_diff_le sum_distrib_right sum_distrib_left)
finally show ?case .
qed simp
corollary power_diff_sumr2: \<comment>\<open>\<open>COMPLEX_POLYFUN\<close> in HOL Light\<close>
fixes x :: "'a::{comm_ring,monoid_mult}"
shows "x^n - y^n = (x - y) * (\<Sum>i<n. y^(n - Suc i) * x^i)"
-using diff_power_eq_setsum[of x "n - 1" y]
+using diff_power_eq_sum[of x "n - 1" y]
by (cases "n = 0") (simp_all add: field_simps)
lemma power_diff_1_eq:
fixes x :: "'a::{comm_ring,monoid_mult}"
shows "n \<noteq> 0 \<Longrightarrow> x^n - 1 = (x - 1) * (\<Sum>i<n. (x^i))"
-using diff_power_eq_setsum [of x _ 1]
+using diff_power_eq_sum [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<n. x^(n - Suc i))"
-using diff_power_eq_setsum [of 1 _ x]
+using diff_power_eq_sum [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<n. x^i)"
-by (metis one_diff_power_eq' [of n x] nat_diff_setsum_reindex)
+by (metis one_diff_power_eq' [of n x] nat_diff_sum_reindex)
subsubsection \<open>The formula for arithmetic sums\<close>
@@ -1952,11 +1952,11 @@
have
"(\<Sum>i\<in>{..<n}. a+?I i*d) =
((\<Sum>i\<in>{..<n}. a) + (\<Sum>i\<in>{..<n}. ?I i*d))"
- by (rule setsum.distrib)
+ by (rule sum.distrib)
also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
unfolding One_nat_def
- by (simp add: setsum_distrib_left atLeast0LessThan[symmetric] setsum_shift_lb_Suc0_0_upt ac_simps)
+ by (simp add: sum_distrib_left atLeast0LessThan[symmetric] sum_shift_lb_Suc0_0_upt ac_simps)
also have "2*\<dots> = 2*?n*a + d*2*(\<Sum>i\<in>{1..<n}. ?I i)"
by (simp add: algebra_simps)
also from ngt1 have "{1..<n} = {1..n - 1}"
@@ -1989,7 +1989,7 @@
by (fact arith_series_general) (* FIXME: duplicate *)
lemma sum_diff_distrib: "\<forall>x. Q x \<le> P x \<Longrightarrow> (\<Sum>x<n. P x) - (\<Sum>x<n. Q x) = (\<Sum>x<n. P x - Q x :: nat)"
- by (subst setsum_subtractf_nat) auto
+ by (subst sum_subtractf_nat) auto
subsubsection \<open>Division remainder\<close>
@@ -2124,13 +2124,13 @@
qed
qed
-lemma setsum_atLeastAtMost_code:
- "setsum f {a..b} = fold_atLeastAtMost_nat (\<lambda>a acc. f a + acc) a b 0"
+lemma sum_atLeastAtMost_code:
+ "sum f {a..b} = fold_atLeastAtMost_nat (\<lambda>a acc. f a + acc) a b 0"
proof -
have "comp_fun_commute (\<lambda>a. op + (f a))"
by unfold_locales (auto simp: o_def add_ac)
thus ?thesis
- by (simp add: setsum.eq_fold fold_atLeastAtMost_nat o_def)
+ by (simp add: sum.eq_fold fold_atLeastAtMost_nat o_def)
qed
lemma setprod_atLeastAtMost_code:
--- a/src/HOL/Transcendental.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/Transcendental.thy Mon Oct 17 14:37:32 2016 +0200
@@ -193,14 +193,14 @@
also have "{0..2*n} = {0<..<2*n} \<union> {0,2*n}" by auto
also have "(\<Sum>k\<in>\<dots>. (2*n) choose k) =
(\<Sum>k\<in>{0<..<2*n}. (2*n) choose k) + (\<Sum>k\<in>{0,2*n}. (2*n) choose k)"
- by (subst setsum.union_disjoint) auto
+ by (subst sum.union_disjoint) auto
also have "(\<Sum>k\<in>{0,2*n}. (2*n) choose k) \<le> (\<Sum>k\<le>1. (n choose k)\<^sup>2)"
by (cases n) simp_all
also from assms have "\<dots> \<le> (\<Sum>k\<le>n. (n choose k)\<^sup>2)"
- by (intro setsum_mono3) auto
+ by (intro sum_mono3) auto
also have "\<dots> = (2*n) choose n" by (rule choose_square_sum)
also have "(\<Sum>k\<in>{0<..<2*n}. (2*n) choose k) \<le> (\<Sum>k\<in>{0<..<2*n}. (2*n) choose n)"
- by (intro setsum_mono binomial_maximum')
+ by (intro sum_mono binomial_maximum')
also have "\<dots> = card {0<..<2*n} * ((2*n) choose n)" by simp
also have "card {0<..<2*n} \<le> 2*n - 1" by (cases n) simp_all
also have "(2 * n - 1) * (2 * n choose n) + (2 * n choose n) = ((2*n) choose n) * (2*n)"
@@ -353,14 +353,14 @@
fix r :: real
assume "0 < r"
from \<open>g sums x\<close>[unfolded sums_def, THEN LIMSEQ_D, OF this]
- obtain no where no_eq: "\<And>n. n \<ge> no \<Longrightarrow> (norm (setsum g {..<n} - x) < r)"
+ obtain no where no_eq: "\<And>n. n \<ge> no \<Longrightarrow> (norm (sum g {..<n} - x) < r)"
by blast
let ?SUM = "\<lambda> m. \<Sum>i<m. if even i then 0 else g ((i - 1) div 2)"
have "(norm (?SUM m - x) < r)" if "m \<ge> 2 * no" for m
proof -
from that have "m div 2 \<ge> no" by auto
- have sum_eq: "?SUM (2 * (m div 2)) = setsum g {..< m div 2}"
+ have sum_eq: "?SUM (2 * (m div 2)) = sum g {..< m div 2}"
using sum_split_even_odd by auto
then have "(norm (?SUM (2 * (m div 2)) - x) < r)"
using no_eq unfolding sum_eq using \<open>m div 2 \<ge> no\<close> by auto
@@ -400,7 +400,7 @@
have "?s sums y" using sums_if'[OF \<open>f sums y\<close>] .
from this[unfolded sums_def, THEN LIMSEQ_Suc]
have "(\<lambda>n. if even n then f (n div 2) else 0) sums y"
- by (simp add: lessThan_Suc_eq_insert_0 setsum_atLeast1_atMost_eq image_Suc_lessThan
+ by (simp add: lessThan_Suc_eq_insert_0 sum_atLeast1_atMost_eq image_Suc_lessThan
if_eq sums_def cong del: if_weak_cong)
from sums_add[OF g_sums this] show ?thesis
by (simp only: if_sum)
@@ -580,10 +580,10 @@
have ?pos using \<open>0 \<le> ?a 0\<close> by auto
moreover have ?neg
using leibniz(2,4)
- unfolding mult_minus_right setsum_negf move_minus neg_le_iff_le
+ unfolding mult_minus_right sum_negf move_minus neg_le_iff_le
by auto
moreover have ?f and ?g
- using leibniz(3,5)[unfolded mult_minus_right setsum_negf move_minus, THEN tendsto_minus_cancel]
+ using leibniz(3,5)[unfolded mult_minus_right sum_negf move_minus, THEN tendsto_minus_cancel]
by auto
ultimately show ?thesis by auto
qed
@@ -614,13 +614,13 @@
(\<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 {..<n} - of_nat n * r = (\<Sum>i<n. f i - r)"
+lemma sumr_diff_mult_const2: "sum f {..<n} - of_nat n * r = (\<Sum>i<n. f i - r)"
for r :: "'a::ring_1"
- by (simp add: setsum_subtractf)
+ by (simp add: sum_subtractf)
lemma lemma_realpow_rev_sumr:
"(\<Sum>p<Suc n. (x ^ p) * (y ^ (n - p))) = (\<Sum>p<Suc n. (x ^ (n - p)) * (y ^ p))"
- by (subst nat_diff_setsum_reindex[symmetric]) simp
+ by (subst nat_diff_sum_reindex[symmetric]) simp
lemma lemma_termdiff2:
fixes h :: "'a::field"
@@ -634,27 +634,27 @@
apply (simp add: mult.assoc [symmetric])
apply (cases n)
apply simp
- apply (simp add: diff_power_eq_setsum h right_diff_distrib [symmetric] mult.assoc
- del: power_Suc setsum_lessThan_Suc of_nat_Suc)
+ apply (simp add: diff_power_eq_sum h right_diff_distrib [symmetric] mult.assoc
+ del: power_Suc sum_lessThan_Suc of_nat_Suc)
apply (subst lemma_realpow_rev_sumr)
apply (subst sumr_diff_mult_const2)
apply simp
- apply (simp only: lemma_termdiff1 setsum_distrib_left)
- apply (rule setsum.cong [OF refl])
+ apply (simp only: lemma_termdiff1 sum_distrib_left)
+ apply (rule sum.cong [OF refl])
apply (simp add: less_iff_Suc_add)
apply clarify
- apply (simp add: setsum_distrib_left diff_power_eq_setsum ac_simps
- del: setsum_lessThan_Suc power_Suc)
+ apply (simp add: sum_distrib_left diff_power_eq_sum ac_simps
+ del: sum_lessThan_Suc power_Suc)
apply (subst mult.assoc [symmetric], subst power_add [symmetric])
apply (simp add: ac_simps)
done
-lemma real_setsum_nat_ivl_bounded2:
+lemma real_sum_nat_ivl_bounded2:
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 {..<n-k} \<le> of_nat n * K"
- apply (rule order_trans [OF setsum_mono])
+ shows "sum f {..<n-k} \<le> of_nat n * K"
+ apply (rule order_trans [OF sum_mono])
apply (rule f)
apply simp
apply (simp add: mult_right_mono K)
@@ -683,8 +683,8 @@
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]
- real_setsum_nat_ivl_bounded2
+ order_trans [OF norm_sum]
+ real_sum_nat_ivl_bounded2
mult_nonneg_nonneg
of_nat_0_le_iff
zero_le_power K)
@@ -1123,7 +1123,7 @@
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)
+ proof (rule sum_strict_mono)
fix n
assume "n \<in> {..< ?N}"
have "\<bar>x\<bar> < S" using \<open>\<bar>x\<bar> < S\<close> .
@@ -1145,7 +1145,7 @@
by blast
qed auto
also have "\<dots> = of_nat (card {..<?N}) * ?r"
- by (rule setsum_constant)
+ by (rule sum_constant)
also have "\<dots> = real ?N * ?r"
by simp
also have "\<dots> = r/3"
@@ -1209,14 +1209,14 @@
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<Suc n. x ^ p * y ^ (n - p)\<bar>"
- unfolding right_diff_distrib[symmetric] diff_power_eq_setsum abs_mult
+ unfolding right_diff_distrib[symmetric] diff_power_eq_sum 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<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)"
- by (rule setsum_abs)
+ by (rule sum_abs)
also have "\<dots> \<le> (\<Sum>p<Suc n. R' ^ n)"
- proof (rule setsum_mono)
+ proof (rule sum_mono)
fix p
assume "p \<in> {..<Suc n}"
then have "p \<le> n" by auto
@@ -1234,7 +1234,7 @@
unfolding power_add[symmetric] using \<open>p \<le> n\<close> by auto
qed
also have "\<dots> = real (Suc n) * R' ^ n"
- unfolding setsum_constant card_atLeastLessThan by auto
+ unfolding sum_constant card_atLeastLessThan by auto
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_of_nonneg[OF zero_le_power[OF less_imp_le[OF \<open>0 < R'\<close>]]]
by linarith
@@ -1448,7 +1448,7 @@
also have "\<dots> = x * (\<Sum>i\<le>n. S x i * S y (n - i)) + y * (\<Sum>i\<le>n. S x i * S y (n - i))"
by (rule distrib_right)
also have "\<dots> = (\<Sum>i\<le>n. x * S x i * S y (n - i)) + (\<Sum>i\<le>n. S x i * y * S y (n - i))"
- by (simp add: setsum_distrib_left ac_simps S_comm)
+ by (simp add: sum_distrib_left ac_simps S_comm)
also have "\<dots> = (\<Sum>i\<le>n. x * S x i * S y (n - i)) + (\<Sum>i\<le>n. S x i * (y * S y (n - i)))"
by (simp add: ac_simps)
also have "\<dots> = (\<Sum>i\<le>n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) +
@@ -1456,19 +1456,19 @@
by (simp add: times_S Suc_diff_le)
also have "(\<Sum>i\<le>n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) =
(\<Sum>i\<le>Suc n. real i *\<^sub>R (S x i * S y (Suc n - i)))"
- by (subst setsum_atMost_Suc_shift) simp
+ by (subst sum_atMost_Suc_shift) simp
also have "(\<Sum>i\<le>n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) =
(\<Sum>i\<le>Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))"
by simp
also have "(\<Sum>i\<le>Suc n. real i *\<^sub>R (S x i * S y (Suc n - i))) +
(\<Sum>i\<le>Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) =
(\<Sum>i\<le>Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n - i)))"
- by (simp only: setsum.distrib [symmetric] scaleR_left_distrib [symmetric]
+ by (simp only: sum.distrib [symmetric] scaleR_left_distrib [symmetric]
of_nat_add [symmetric]) simp
also have "\<dots> = real (Suc n) *\<^sub>R (\<Sum>i\<le>Suc n. S x i * S y (Suc n - i))"
- by (simp only: scaleR_right.setsum)
+ by (simp only: scaleR_right.sum)
finally show "S (x + y) (Suc n) = (\<Sum>i\<le>Suc n. S x i * S y (Suc n - i))"
- by (simp del: setsum_cl_ivl_Suc)
+ by (simp del: sum_cl_ivl_Suc)
qed
lemma exp_add_commuting: "x * y = y * x \<Longrightarrow> exp (x + y) = exp x * exp y"
@@ -1522,7 +1522,7 @@
corollary exp_real_of_nat_mult: "exp (real n * x) = exp x ^ n"
by (simp add: exp_of_nat_mult)
-lemma exp_setsum: "finite I \<Longrightarrow> exp (setsum f I) = setprod (\<lambda>x. exp (f x)) I"
+lemma exp_sum: "finite I \<Longrightarrow> exp (sum f I) = setprod (\<lambda>x. exp (f x)) I"
by (induct I rule: finite_induct) (auto simp: exp_add_commuting mult.commute)
lemma exp_divide_power_eq:
@@ -1596,7 +1596,7 @@
have "1 + x \<le> (\<Sum>n<2. inverse (fact n) * x^n)"
by (auto simp add: numeral_2_eq_2)
also have "\<dots> \<le> (\<Sum>n. inverse (fact n) * x^n)"
- apply (rule setsum_le_suminf [OF summable_exp])
+ apply (rule sum_le_suminf [OF summable_exp])
using \<open>0 < x\<close>
apply (auto simp add: zero_le_mult_iff)
done
@@ -1744,7 +1744,7 @@
for x :: real
by (rule ln_unique) (simp add: exp_add)
-lemma ln_setprod: "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i > 0) \<Longrightarrow> ln (setprod f I) = setsum (\<lambda>x. ln(f x)) I"
+lemma ln_setprod: "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i > 0) \<Longrightarrow> ln (setprod f I) = sum (\<lambda>x. ln(f x)) I"
for f :: "'a \<Rightarrow> real"
by (induct I rule: finite_induct) (auto simp: ln_mult setprod_pos)
@@ -3340,7 +3340,7 @@
(if even p
then of_real ((-1) ^ (p div 2) / (fact p)) * (\<Sum>n\<le>p. (p choose n) *\<^sub>R (x^n) * y^(p-n))
else 0)"
- by (auto simp: setsum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide)
+ by (auto simp: sum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide)
also have "\<dots> = cos_coeff p *\<^sub>R ((x + y) ^ p)"
by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost)
finally show ?thesis .
@@ -3372,7 +3372,7 @@
"(\<lambda>p. \<Sum>n\<le>p. (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0))
sums (cos x * cos y - sin x * sin y)"
using sums_diff [OF cos_x_cos_y [of x y] sin_x_sin_y [of x y]]
- by (simp add: setsum_subtractf [symmetric])
+ by (simp add: sum_subtractf [symmetric])
then show ?thesis
by (blast intro: sums_cos_x_plus_y sums_unique2)
qed
@@ -3531,7 +3531,7 @@
fixes f :: "nat \<Rightarrow> real"
shows "summable f \<Longrightarrow>
(\<And>d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc (Suc 0) * d) + 1))) \<Longrightarrow>
- setsum f {..<k} < suminf f"
+ sum f {..<k} < suminf f"
apply (simp only: One_nat_def)
apply (subst suminf_split_initial_segment [where k=k])
apply assumption
@@ -5629,7 +5629,7 @@
moreover have "isCont (\<lambda> x. ?a x n - ?diff x n) x" for x
unfolding diff_conv_add_uminus divide_inverse
by (auto intro!: isCont_add isCont_rabs continuous_ident isCont_minus isCont_arctan
- isCont_inverse isCont_mult isCont_power continuous_const isCont_setsum
+ isCont_inverse isCont_mult isCont_power continuous_const isCont_sum
simp del: add_uminus_conv_diff)
ultimately have "0 \<le> ?a 1 n - ?diff 1 n"
by (rule LIM_less_bound)
@@ -5804,8 +5804,8 @@
for m :: nat
by auto
-lemma setsum_up_index_split: "(\<Sum>k\<le>m + n. f k) = (\<Sum>k\<le>m. f k) + (\<Sum>k = Suc m..m + n. f k)"
- by (metis atLeast0AtMost Suc_eq_plus1 le0 setsum_ub_add_nat)
+lemma sum_up_index_split: "(\<Sum>k\<le>m + n. f k) = (\<Sum>k\<le>m. f k) + (\<Sum>k = Suc m..m + n. f k)"
+ by (metis atLeast0AtMost Suc_eq_plus1 le0 sum_ub_add_nat)
lemma Sigma_interval_disjoint: "(SIGMA i:A. {..v i}) \<inter> (SIGMA i:A.{v i<..w}) = {}"
for w :: "'a::order"
@@ -5823,19 +5823,19 @@
(\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
proof -
have "(\<Sum>i\<le>m. (a i) * x ^ i) * (\<Sum>j\<le>n. (b j) * x ^ j) = (\<Sum>i\<le>m. \<Sum>j\<le>n. (a i * x ^ i) * (b j * x ^ j))"
- by (rule setsum_product)
+ by (rule sum_product)
also have "\<dots> = (\<Sum>i\<le>m + n. \<Sum>j\<le>n + m. a i * x ^ i * (b j * x ^ j))"
- using assms by (auto simp: setsum_up_index_split)
+ using assms by (auto simp: sum_up_index_split)
also have "\<dots> = (\<Sum>r\<le>m + n. \<Sum>j\<le>m + n - r. a r * x ^ r * (b j * x ^ j))"
- apply (simp add: add_ac setsum.Sigma product_atMost_eq_Un)
- apply (clarsimp simp add: setsum_Un Sigma_interval_disjoint intro!: setsum.neutral)
+ apply (simp add: add_ac sum.Sigma product_atMost_eq_Un)
+ apply (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral)
apply (metis add_diff_assoc2 add.commute add_lessD1 leD m n nat_le_linear neqE)
done
also have "\<dots> = (\<Sum>(i,j)\<in>{(i,j). i+j \<le> m+n}. (a i * x ^ i) * (b j * x ^ j))"
- by (auto simp: pairs_le_eq_Sigma setsum.Sigma)
+ by (auto simp: pairs_le_eq_Sigma sum.Sigma)
also have "\<dots> = (\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
- apply (subst setsum_triangle_reindex_eq)
- apply (auto simp: algebra_simps setsum_distrib_left intro!: setsum.cong)
+ apply (subst sum_triangle_reindex_eq)
+ apply (auto simp: algebra_simps sum_distrib_left intro!: sum.cong)
apply (metis le_add_diff_inverse power_add)
done
finally show ?thesis .
@@ -5849,7 +5849,7 @@
(\<Sum>r\<le>m + n. (\<Sum>k\<le>r. (a k) * (b (r - k))) * x ^ r)"
using polynomial_product [of m a n b x] assms
by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric]
- of_nat_eq_iff Int.int_setsum [symmetric])
+ of_nat_eq_iff Int.int_sum [symmetric])
lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*)
fixes x :: "'a::idom"
@@ -5860,19 +5860,19 @@
have h: "bij_betw (\<lambda>(i,j). (j,i)) ((SIGMA i : atMost n. lessThan i)) (SIGMA j : lessThan n. {Suc j..n})"
by (auto simp: bij_betw_def inj_on_def)
have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) = (\<Sum>i\<le>n. a i * (x^i - y^i))"
- by (simp add: right_diff_distrib setsum_subtractf)
+ by (simp add: right_diff_distrib sum_subtractf)
also have "\<dots> = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
by (simp add: power_diff_sumr2 mult.assoc)
also have "\<dots> = (\<Sum>i\<le>n. \<Sum>j<i. a i * (x - y) * (y^(i - Suc j) * x^j))"
- by (simp add: setsum_distrib_left)
+ by (simp add: sum_distrib_left)
also have "\<dots> = (\<Sum>(i,j) \<in> (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))"
- by (simp add: setsum.Sigma)
+ by (simp add: sum.Sigma)
also have "\<dots> = (\<Sum>(j,i) \<in> (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))"
- by (auto simp add: setsum.reindex_bij_betw [OF h, symmetric] intro: setsum.strong_cong)
+ by (auto simp add: sum.reindex_bij_betw [OF h, symmetric] intro: sum.strong_cong)
also have "\<dots> = (\<Sum>j<n. \<Sum>i=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))"
- by (simp add: setsum.Sigma)
+ by (simp add: sum.Sigma)
also have "\<dots> = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - j - 1)) * x^j)"
- by (simp add: setsum_distrib_left mult_ac)
+ by (simp add: sum_distrib_left mult_ac)
finally show ?thesis .
qed
@@ -5891,10 +5891,10 @@
apply (auto simp: )
done
then show ?thesis
- by (auto simp add: setsum.reindex_bij_betw [OF h, symmetric] intro: setsum.strong_cong)
+ by (auto simp add: sum.reindex_bij_betw [OF h, symmetric] intro: sum.strong_cong)
qed
then show ?thesis
- by (simp add: polyfun_diff [OF assms] setsum_distrib_right)
+ by (simp add: polyfun_diff [OF assms] sum_distrib_right)
qed
lemma polyfun_linear_factor: (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*)
@@ -5946,10 +5946,10 @@
have "(\<Sum>i\<le>Suc n. c i * w^i) = w * (\<Sum>i\<le>n. c (Suc i) * w^i)" for w
proof -
have "(\<Sum>i\<le>Suc n. c i * w^i) = (\<Sum>i\<le>n. c (Suc i) * w ^ Suc i)"
- unfolding Set_Interval.setsum_atMost_Suc_shift
+ unfolding Set_Interval.sum_atMost_Suc_shift
by simp
also have "\<dots> = w * (\<Sum>i\<le>n. c (Suc i) * w^i)"
- by (simp add: setsum_distrib_left ac_simps)
+ by (simp add: sum_distrib_left ac_simps)
finally show ?thesis .
qed
then have w: "\<And>w. w \<noteq> 0 \<Longrightarrow> (\<Sum>i\<le>n. c (Suc i) * w^i) = 0"
@@ -6008,7 +6008,7 @@
by blast
show ?succase
using Suc.IH [of b k'] bk'
- by (simp add: eq card_insert_if del: setsum_atMost_Suc)
+ by (simp add: eq card_insert_if del: sum_atMost_Suc)
qed
qed
@@ -6048,7 +6048,7 @@
for c :: "nat \<Rightarrow> 'a::{idom,real_normed_div_algebra}"
proof -
have "(\<forall>x. (\<Sum>i\<le>n. c i * x^i) = (\<Sum>i\<le>n. d i * x^i)) \<longleftrightarrow> (\<forall>x. (\<Sum>i\<le>n. (c i - d i) * x^i) = 0)"
- by (simp add: left_diff_distrib Groups_Big.setsum_subtractf)
+ by (simp add: left_diff_distrib Groups_Big.sum_subtractf)
also have "\<dots> \<longleftrightarrow> (\<forall>i\<le>n. c i - d i = 0)"
by (rule polyfun_eq_0)
finally show ?thesis
@@ -6078,7 +6078,7 @@
fixes z :: "'a::idom"
assumes "1 \<le> n"
shows "z^n = a \<longleftrightarrow> (\<Sum>i\<le>n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0"
- using assms by (cases n) (simp_all add: setsum_head_Suc atLeast0AtMost [symmetric])
+ using assms by (cases n) (simp_all add: sum_head_Suc atLeast0AtMost [symmetric])
lemma
assumes "SORT_CONSTRAINT('a::{idom,real_normed_div_algebra})"
--- a/src/HOL/UNITY/Comp/Alloc.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/UNITY/Comp/Alloc.thy Mon Oct 17 14:37:32 2016 +0200
@@ -903,9 +903,9 @@
@{thm Always_tokens_giv_le_allocGiv}, @{thm Always_tokens_allocRel_le_rel}] RS
@{thm Always_weaken}] 1\<close>)
apply auto
- apply (rule setsum_fun_mono [THEN order_trans])
+ apply (rule sum_fun_mono [THEN order_trans])
apply (drule_tac [2] order_trans)
- apply (rule_tac [2] add_le_mono [OF order_refl setsum_fun_mono])
+ apply (rule_tac [2] add_le_mono [OF order_refl sum_fun_mono])
prefer 3 apply assumption
apply auto
done
--- a/src/HOL/UNITY/Comp/AllocBase.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/UNITY/Comp/AllocBase.thy Mon Oct 17 14:37:32 2016 +0200
@@ -19,9 +19,9 @@
abbreviation (input)
"bag_of \<equiv> mset"
-lemma setsum_fun_mono:
+lemma sum_fun_mono:
fixes f :: "nat \<Rightarrow> nat"
- shows "(\<And>i. i < n \<Longrightarrow> f i \<le> g i) \<Longrightarrow> setsum f {..<n} \<le> setsum g {..<n}"
+ shows "(\<And>i. i < n \<Longrightarrow> f i \<le> g i) \<Longrightarrow> sum f {..<n} \<le> sum g {..<n}"
by (induct n) (auto simp add: lessThan_Suc add_le_mono)
lemma tokens_mono_prefix: "xs \<le> ys \<Longrightarrow> tokens xs \<le> tokens ys"
@@ -44,14 +44,14 @@
apply simp
done
-(** setsum **)
+(** sum **)
-declare setsum.cong [cong]
+declare sum.cong [cong]
lemma bag_of_sublist_lemma:
"(\<Sum>i\<in> A Int lessThan k. {#if i<k then f i else g i#}) =
(\<Sum>i\<in> A Int lessThan k. {#f i#})"
-by (rule setsum.cong, auto)
+by (rule sum.cong, auto)
lemma bag_of_sublist:
"bag_of (sublist l A) =
@@ -67,7 +67,7 @@
bag_of (sublist l A) + bag_of (sublist l B)"
apply (subgoal_tac "A Int B Int {..<length l} =
(A Int {..<length l}) Int (B Int {..<length l}) ")
-apply (simp add: bag_of_sublist Int_Un_distrib2 setsum.union_inter, blast)
+apply (simp add: bag_of_sublist Int_Un_distrib2 sum.union_inter, blast)
done
lemma bag_of_sublist_Un_disjoint:
@@ -82,7 +82,7 @@
(\<Sum>i\<in>I. bag_of (sublist l (A i)))"
apply (auto simp add: bag_of_sublist)
unfolding UN_simps [symmetric]
-apply (subst setsum.UNION_disjoint)
+apply (subst sum.UNION_disjoint)
apply auto
done
--- a/src/HOL/UNITY/Comp/AllocImpl.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/UNITY/Comp/AllocImpl.thy Mon Oct 17 14:37:32 2016 +0200
@@ -296,7 +296,7 @@
(bag_of o merge.Out) Fols
(%s. \<Sum>i \<in> lessThan Nclients. (bag_of o sub i o merge.In) s)"
apply (rule Merge_Bag_Follows_lemma [THEN Always_Follows1, THEN guaranteesI], auto)
-apply (rule Follows_setsum)
+apply (rule Follows_sum)
apply (cut_tac Merge_spec)
apply (auto simp add: merge_spec_def merge_follows_def o_def)
apply (drule guaranteesD)
@@ -360,7 +360,7 @@
(%s. bag_of (sublist (distr.In s) (lessThan (length(distr.iIn s))))))"
apply (rule guaranteesI, clarify)
apply (rule Distr_Bag_Follows_lemma [THEN Always_Follows2], auto)
-apply (rule Follows_setsum)
+apply (rule Follows_sum)
apply (cut_tac Distrib_spec)
apply (auto simp add: distr_spec_def distr_follows_def o_def)
apply (drule guaranteesD)
--- a/src/HOL/UNITY/Follows.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/UNITY/Follows.thy Mon Oct 17 14:37:32 2016 +0200
@@ -236,7 +236,7 @@
apply (blast intro: Follows_union_lemma)
done
-lemma Follows_setsum:
+lemma Follows_sum:
"!!f ::['c,'b] => ('a::order) multiset.
[| \<forall>i \<in> I. F \<in> f' i Fols f i; finite I |]
==> F \<in> (%s. \<Sum>i \<in> I. f' i s) Fols (%s. \<Sum>i \<in> I. f i s)"
--- a/src/HOL/ex/Execute_Choice.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/Execute_Choice.thy Mon Oct 17 14:37:32 2016 +0200
@@ -54,7 +54,7 @@
"finite (Mapping.keys M) \<Longrightarrow> x \<in> Mapping.keys M \<Longrightarrow> y \<in> Mapping.keys M \<Longrightarrow>
the (Mapping.lookup M x) + valuesum (Mapping.delete x M) =
the (Mapping.lookup M y) + valuesum (Mapping.delete y M)"
- unfolding valuesum_def by transfer (simp add: setsum_diff)
+ unfolding valuesum_def by transfer (simp add: sum_diff)
text \<open>
Given \<open>valuesum_rec\<close> as initial description, we stepwise refine it to something executable;
--- a/src/HOL/ex/HarmonicSeries.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/HarmonicSeries.thy Mon Oct 17 14:37:32 2016 +0200
@@ -87,7 +87,7 @@
qed
then have
"(\<Sum>n\<in>(?S m). 1 / real n) \<ge> (\<Sum>n\<in>(?S m). 1/(real tm))"
- by (rule setsum_mono)
+ by (rule sum_mono)
moreover have
"(\<Sum>n\<in>(?S m). 1/(real tm)) = 1/2"
proof -
@@ -147,7 +147,7 @@
"\<dots> = (\<Sum>n\<in>{(1::nat)..2}. 1/real n)" by simp
also have
"\<dots> = ((\<Sum>n\<in>{Suc 1..2}. 1/real n) + 1/(real (1::nat)))"
- by (subst setsum_head)
+ by (subst sum_head)
(auto simp: atLeastSucAtMost_greaterThanAtMost)
also have
"\<dots> = ((\<Sum>n\<in>{2..2::nat}. 1/real n) + 1/(real (1::nat)))"
@@ -194,7 +194,7 @@
"finite {1..(2::nat)^M}" and "finite {(2::nat)^M+1..(2::nat)^(Suc M)}"
by auto
ultimately show ?thesis
- by (auto intro: setsum.union_disjoint)
+ by (auto intro: sum.union_disjoint)
qed
moreover
{
@@ -243,7 +243,7 @@
let ?f = "(\<lambda>x. 1/2)"
let ?g = "(\<lambda>x. (\<Sum>n\<in>{(2::nat)^(x - 1)+1..2^x}. 1/real n))"
from harmonic_aux have "\<And>x. x\<in>{1..M} \<Longrightarrow> ?f x \<le> ?g x" by simp
- then have "(\<Sum>m\<in>{1..M}. ?g m) \<ge> (\<Sum>m\<in>{1..M}. ?f m)" by (rule setsum_mono)
+ then have "(\<Sum>m\<in>{1..M}. ?g m) \<ge> (\<Sum>m\<in>{1..M}. ?f m)" by (rule sum_mono)
thus ?thesis by simp
qed
finally have "(?P M) \<ge> (1 + (\<Sum>m\<in>{1..M}. 1/2))" .
@@ -266,8 +266,8 @@
text \<open>The final theorem shows that as we take more and more elements
(see @{thm [source] harmonic_aux3}) we get an ever increasing sum. By assuming
-the sum converges, the lemma @{thm [source] setsum_less_suminf} ( @{thm
-setsum_less_suminf} ) states that each sum is bounded above by the
+the sum converges, the lemma @{thm [source] sum_less_suminf} ( @{thm
+sum_less_suminf} ) states that each sum is bounded above by the
series' limit. This contradicts our first statement and thus we prove
that the harmonic series is divergent.\<close>
@@ -281,9 +281,9 @@
then have ngt: "1 + real n/2 > ?s" by linarith
define j where "j = (2::nat)^n"
have "\<forall>m\<ge>j. 0 < ?f m" by simp
- with sf have "(\<Sum>i<j. ?f i) < ?s" by (rule setsum_less_suminf)
+ with sf have "(\<Sum>i<j. ?f i) < ?s" by (rule sum_less_suminf)
then have "(\<Sum>i\<in>{Suc 0..<Suc j}. 1/(real i)) < ?s"
- unfolding setsum_shift_bounds_Suc_ivl by (simp add: atLeast0LessThan)
+ unfolding sum_shift_bounds_Suc_ivl by (simp add: atLeast0LessThan)
with j_def have
"(\<Sum>i\<in>{1..< Suc ((2::nat)^n)}. 1 / (real i)) < ?s" by simp
then have
--- a/src/HOL/ex/Meson_Test.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/Meson_Test.thy Mon Oct 17 14:37:32 2016 +0200
@@ -10,7 +10,7 @@
below and constants declared in HOL!
\<close>
-hide_const (open) implies union inter subset quotient
+hide_const (open) implies union inter subset quotient sum
text \<open>
Test data for the MESON proof procedure
--- a/src/HOL/ex/Sum_of_Powers.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/Sum_of_Powers.thy Mon Oct 17 14:37:32 2016 +0200
@@ -55,14 +55,14 @@
qed (simp add: assms)
qed
-lemma setsum_diff: "((\<Sum>i\<le>n::nat. f (i + 1) - f i)::'a::field) = f (n + 1) - f 0"
+lemma sum_diff: "((\<Sum>i\<le>n::nat. f (i + 1) - f i)::'a::field) = f (n + 1) - f 0"
by (induct n) (auto simp add: field_simps)
declare One_nat_def [simp del]
subsection \<open>Bernoulli Numbers and Bernoulli Polynomials\<close>
-declare setsum.cong [fundef_cong]
+declare sum.cong [fundef_cong]
fun bernoulli :: "nat \<Rightarrow> real"
where
@@ -84,12 +84,12 @@
next
case (Suc n')
have "(\<Sum>k\<le>n'. real (Suc n' choose k) * bernoulli k * 0 ^ (Suc n' - k)) = 0"
- by (rule setsum.neutral) auto
+ by (rule sum.neutral) auto
with Suc show ?thesis
unfolding bernpoly_def by simp
qed
-lemma setsum_binomial_times_bernoulli:
+lemma sum_binomial_times_bernoulli:
"(\<Sum>k\<le>n. ((Suc n) choose k) * bernoulli k) = (if n = 0 then 1 else 0)"
proof (cases n)
case 0
@@ -110,7 +110,7 @@
unfolding bernpoly_def by (rule DERIV_cong) (fast intro!: derivative_intros, simp)
moreover have "(\<Sum>k\<le>n. real (Suc n - k) * x ^ (n - k) * (real (Suc n choose k) * bernoulli k)) = (n + 1) * bernpoly n x"
unfolding bernpoly_def
- by (auto intro: setsum.cong simp add: setsum_distrib_left real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff)
+ by (auto intro: sum.cong simp add: sum_distrib_left real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff)
ultimately show ?thesis by auto
qed
@@ -122,7 +122,7 @@
next
case (Suc n)
have "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = (Suc n) * 0 ^ n"
- unfolding bernpoly_0 unfolding bernpoly_def by (simp add: setsum_binomial_times_bernoulli zero_power)
+ unfolding bernpoly_0 unfolding bernpoly_def by (simp add: sum_binomial_times_bernoulli zero_power)
then have const: "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = real (Suc n) * 0 ^ n" by (simp add: power_0_left)
have hyps': "\<And>x. (real n + 1) * bernpoly n (x + 1) - (real n + 1) * bernpoly n x = real n * x ^ (n - Suc 0) * real (Suc n)"
unfolding right_diff_distrib[symmetric] by (simp add: Suc.hyps One_nat_def)
@@ -135,11 +135,11 @@
lemma sum_of_powers: "(\<Sum>k\<le>n::nat. (real k) ^ m) = (bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0) / (m + 1)"
proof -
from diff_bernpoly[of "Suc m", simplified] have "(m + (1::real)) * (\<Sum>k\<le>n. (real k) ^ m) = (\<Sum>k\<le>n. bernpoly (Suc m) (real k + 1) - bernpoly (Suc m) (real k))"
- by (auto simp add: setsum_distrib_left intro!: setsum.cong)
+ by (auto simp add: sum_distrib_left intro!: sum.cong)
also have "... = (\<Sum>k\<le>n. bernpoly (Suc m) (real (k + 1)) - bernpoly (Suc m) (real k))"
by simp
also have "... = bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0"
- by (simp only: setsum_diff[where f="\<lambda>k. bernpoly (Suc m) (real k)"]) simp
+ by (simp only: sum_diff[where f="\<lambda>k. bernpoly (Suc m) (real k)"]) simp
finally show ?thesis by (auto simp add: field_simps intro!: eq_divide_imp)
qed
@@ -149,16 +149,16 @@
"n > 0 \<Longrightarrow> (n choose k) = (if k = 0 then 1 else (n - 1) choose (k - 1) + ((n - 1) choose k))"
by (auto simp add: gr0_conv_Suc)
-lemma setsum_unroll:
+lemma sum_unroll:
"(\<Sum>k\<le>n::nat. f k) = (if n = 0 then f 0 else f n + (\<Sum>k\<le>n - 1. f k))"
-by auto (metis One_nat_def Suc_pred add.commute setsum_atMost_Suc)
+by auto (metis One_nat_def Suc_pred add.commute sum_atMost_Suc)
lemma bernoulli_unroll:
"n > 0 \<Longrightarrow> bernoulli n = - 1 / (real n + 1) * (\<Sum>k\<le>n - 1. real (n + 1 choose k) * bernoulli k)"
by (cases n) (simp add: bernoulli.simps One_nat_def)+
lemmas unroll = binomial_unroll
- bernoulli.simps(1) bernoulli_unroll setsum_unroll bernpoly_def
+ bernoulli.simps(1) bernoulli_unroll sum_unroll bernpoly_def
lemma sum_of_squares: "(\<Sum>k\<le>n::nat. k ^ 2) = (2 * n ^ 3 + 3 * n ^ 2 + n) / 6"
proof -
--- a/src/HOL/ex/ThreeDivides.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/ThreeDivides.thy Mon Oct 17 14:37:32 2016 +0200
@@ -113,7 +113,7 @@
shows "(3 dvd (\<Sum>x<nd. D x * 10^x)) = (3 dvd (\<Sum>x<nd. D x))"
proof
have mono: "(\<Sum>x<nd. D x) \<le> (\<Sum>x<nd. D x * 10^x)"
- by (rule setsum_mono) simp
+ by (rule sum_mono) simp
txt \<open>This lets us form the term
@{term "(\<Sum>x<nd. D x * 10^x) - (\<Sum>x<nd. D x)"}\<close>
@@ -193,17 +193,17 @@
"m = 10*(\<Sum>x<nd. m div 10 div 10^x mod 10 * 10^x) + c" by simp
also have
"\<dots> = (\<Sum>x<nd. m div 10 div 10^x mod 10 * 10^(x+1)) + c"
- by (subst setsum_distrib_left) (simp add: ac_simps)
+ by (subst sum_distrib_left) (simp add: ac_simps)
also have
"\<dots> = (\<Sum>x<nd. m div 10^(Suc x) mod 10 * 10^(Suc x)) + c"
by (simp add: div_mult2_eq[symmetric])
also have
"\<dots> = (\<Sum>x\<in>{Suc 0..<Suc nd}. m div 10^x mod 10 * 10^x) + c"
- by (simp only: setsum_shift_bounds_Suc_ivl)
+ by (simp only: sum_shift_bounds_Suc_ivl)
(simp add: atLeast0LessThan)
also have
"\<dots> = (\<Sum>x<Suc nd. m div 10^x mod 10 * 10^x)"
- by (simp add: atLeast0LessThan[symmetric] setsum_head_upt_Suc cdef)
+ by (simp add: atLeast0LessThan[symmetric] sum_head_upt_Suc cdef)
also note \<open>Suc nd = nlen m\<close>
finally
show "m = (\<Sum>x<nlen m. m div 10^x mod 10 * 10^x)" .
--- a/src/HOL/ex/Transfer_Int_Nat.thy Mon Oct 17 11:07:01 2016 +0200
+++ b/src/HOL/ex/Transfer_Int_Nat.thy Mon Oct 17 14:37:32 2016 +0200
@@ -104,12 +104,12 @@
unfolding rel_fun_def ZN_def rel_set_def
by (clarsimp simp add: Bex_def, arith)
-lemma ZN_setsum [transfer_rule]:
- "bi_unique A \<Longrightarrow> ((A ===> ZN) ===> rel_set A ===> ZN) setsum setsum"
+lemma ZN_sum [transfer_rule]:
+ "bi_unique A \<Longrightarrow> ((A ===> ZN) ===> rel_set A ===> ZN) sum sum"
apply (intro rel_funI)
apply (erule (1) bi_unique_rel_set_lemma)
- apply (simp add: setsum.reindex int_setsum ZN_def rel_fun_def)
- apply (rule setsum.cong)
+ apply (simp add: sum.reindex int_sum ZN_def rel_fun_def)
+ apply (rule sum.cong)
apply simp_all
done