--- 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)"