fact consolidation
authorhaftmann
Sat, 28 Jun 2014 09:16:42 +0200
changeset 587606ab1c7cb0b8d
parent 58759 29fe9bac501b
child 58761 2b8b1a8587da
fact consolidation
NEWS
src/HOL/Decision_Procs/Approximation.thy
src/HOL/Deriv.thy
src/HOL/Fact.thy
src/HOL/Groups_Big.thy
src/HOL/Hoare_Parallel/OG_Examples.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy
src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy
src/HOL/Library/BigO.thy
src/HOL/Library/Convex.thy
src/HOL/Library/FSet.thy
src/HOL/Library/Formal_Power_Series.thy
src/HOL/Library/Indicator_Function.thy
src/HOL/Library/Multiset.thy
src/HOL/Library/Permutations.thy
src/HOL/List.thy
src/HOL/MacLaurin.thy
src/HOL/Metis_Examples/Big_O.thy
src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy
src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Determinants.thy
src/HOL/Multivariate_Analysis/Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy
src/HOL/Multivariate_Analysis/Fashoda.thy
src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Linear_Algebra.thy
src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/NSA/HSeries.thy
src/HOL/Nat_Transfer.thy
src/HOL/Number_Theory/Binomial.thy
src/HOL/Number_Theory/Cong.thy
src/HOL/Number_Theory/Gauss.thy
src/HOL/Number_Theory/UniqueFactorization.thy
src/HOL/Old_Number_Theory/Euler.thy
src/HOL/Old_Number_Theory/EulerFermat.thy
src/HOL/Old_Number_Theory/Gauss.thy
src/HOL/Old_Number_Theory/Pocklington.thy
src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy
src/HOL/Old_Number_Theory/WilsonBij.thy
src/HOL/Old_Number_Theory/WilsonRuss.thy
src/HOL/Power.thy
src/HOL/Probability/Bochner_Integration.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Distributions.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Independent_Family.thy
src/HOL/Probability/Infinite_Product_Measure.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Measure_Space.thy
src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy
src/HOL/Probability/Probability_Measure.thy
src/HOL/Probability/Projective_Family.thy
src/HOL/Probability/Projective_Limit.thy
src/HOL/Probability/Sigma_Algebra.thy
src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy
src/HOL/Real_Vector_Spaces.thy
src/HOL/Series.thy
src/HOL/Set_Interval.thy
src/HOL/Transcendental.thy
src/HOL/UNITY/Comp/AllocBase.thy
src/HOL/ex/HarmonicSeries.thy
src/HOL/ex/Transfer_Int_Nat.thy
     1.1 --- a/NEWS	Fri Jun 27 22:08:55 2014 +0200
     1.2 +++ b/NEWS	Sat Jun 28 09:16:42 2014 +0200
     1.3 @@ -375,6 +375,75 @@
     1.4  * Theory reorganizations:
     1.5    * Big_Operators.thy ~> Groups_Big.thy and Lattices_Big.thy
     1.6  
     1.7 +* Consolidated some facts about big group operators:
     1.8 +
     1.9 +    setsum_0' ~> setsum.neutral
    1.10 +    setsum_0 ~> setsum.neutral_const
    1.11 +    setsum_addf ~> setsum.distrib
    1.12 +    setsum_cartesian_product ~> setsum.cartesian_product
    1.13 +    setsum_cases ~> setsum.If_cases
    1.14 +    setsum_commute ~> setsum.commute
    1.15 +    setsum_cong ~> setsum.cong
    1.16 +    setsum_delta ~> setsum.delta
    1.17 +    setsum_delta' ~> setsum.delta'
    1.18 +    setsum_diff1' ~> setsum.remove
    1.19 +    setsum_empty ~> setsum.empty
    1.20 +    setsum_infinite ~> setsum.infinite
    1.21 +    setsum_insert ~> setsum.insert
    1.22 +    setsum_inter_restrict'' ~> setsum.inter_filter
    1.23 +    setsum_mono_zero_cong_left ~> setsum.mono_neutral_cong_left
    1.24 +    setsum_mono_zero_cong_right ~> setsum.mono_neutral_cong_right
    1.25 +    setsum_mono_zero_left ~> setsum.mono_neutral_left
    1.26 +    setsum_mono_zero_right ~> setsum.mono_neutral_right
    1.27 +    setsum_reindex ~> setsum.reindex
    1.28 +    setsum_reindex_cong ~> setsum.reindex_cong
    1.29 +    setsum_reindex_nonzero ~> setsum.reindex_nontrivial
    1.30 +    setsum_restrict_set ~> setsum.inter_restrict
    1.31 +    setsum_Plus ~> setsum.Plus
    1.32 +    setsum_setsum_restrict ~> setsum.commute_restrict
    1.33 +    setsum_Sigma ~> setsum.Sigma
    1.34 +    setsum_subset_diff ~> setsum.subset_diff
    1.35 +    setsum_Un_disjoint ~> setsum.union_disjoint
    1.36 +    setsum_UN_disjoint ~> setsum.UNION_disjoint
    1.37 +    setsum_Un_Int ~> setsum.union_inter
    1.38 +    setsum_Union_disjoint ~> setsum.Union_disjoint
    1.39 +    setsum_UNION_zero ~> setsum.Union_comp
    1.40 +    setsum_Un_zero ~> setsum.union_inter_neutral
    1.41 +    strong_setprod_cong ~> setprod.strong_cong
    1.42 +    strong_setsum_cong ~> setsum.strong_cong
    1.43 +    setprod_1' ~> setprod.neutral
    1.44 +    setprod_1 ~> setprod.neutral_const
    1.45 +    setprod_cartesian_product ~> setprod.cartesian_product
    1.46 +    setprod_cong ~> setprod.cong
    1.47 +    setprod_delta ~> setprod.delta
    1.48 +    setprod_delta' ~> setprod.delta'
    1.49 +    setprod_empty ~> setprod.empty
    1.50 +    setprod_infinite ~> setprod.infinite
    1.51 +    setprod_insert ~> setprod.insert
    1.52 +    setprod_mono_one_cong_left ~> setprod.mono_neutral_cong_left
    1.53 +    setprod_mono_one_cong_right ~> setprod.mono_neutral_cong_right
    1.54 +    setprod_mono_one_left ~> setprod.mono_neutral_left
    1.55 +    setprod_mono_one_right ~> setprod.mono_neutral_right
    1.56 +    setprod_reindex ~> setprod.reindex
    1.57 +    setprod_reindex_cong ~> setprod.reindex_cong
    1.58 +    setprod_reindex_nonzero ~> setprod.reindex_nontrivial
    1.59 +    setprod_Sigma ~> setprod.Sigma
    1.60 +    setprod_subset_diff ~> setprod.subset_diff
    1.61 +    setprod_timesf ~> setprod.distrib
    1.62 +    setprod_Un2 ~> setprod.union_diff2
    1.63 +    setprod_Un_disjoint ~> setprod.union_disjoint
    1.64 +    setprod_UN_disjoint ~> setprod.UNION_disjoint
    1.65 +    setprod_Un_Int ~> setprod.union_inter
    1.66 +    setprod_Union_disjoint ~> setprod.Union_disjoint
    1.67 +    setprod_Un_one ~> setprod.union_inter_neutral
    1.68 +
    1.69 +  Dropped setsum_cong2 (simple variant of setsum.cong).
    1.70 +  Dropped setsum_inter_restrict' (simple variant of setsum.inter_restrict)
    1.71 +  Dropped setsum_reindex_id, setprod_reindex_id
    1.72 +    (simple variants of setsum.reindex [symmetric], setprod.reindex [symmetric]).
    1.73 +
    1.74 +  INCOMPATIBILITY.
    1.75 +
    1.76  * New internal SAT solver "cdclite" that produces models and proof traces.
    1.77    This solver replaces the internal SAT solvers "enumerate" and "dpll".
    1.78    Applications that explicitly used one of these two SAT solvers should
     2.1 --- a/src/HOL/Decision_Procs/Approximation.thy	Fri Jun 27 22:08:55 2014 +0200
     2.2 +++ b/src/HOL/Decision_Procs/Approximation.thy	Sat Jun 28 09:16:42 2014 +0200
     2.3 @@ -34,7 +34,7 @@
     2.4    show ?thesis
     2.5      unfolding setsum_right_distrib shift_pow uminus_add_conv_diff [symmetric] setsum_negf[symmetric]
     2.6      setsum_head_upt_Suc[OF zero_less_Suc]
     2.7 -    setsum_reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n  *a n * x^n"] by auto
     2.8 +    setsum.reindex[OF inj_Suc, unfolded comp_def, symmetric, of "\<lambda> n. (-1)^n  *a n * x^n"] by auto
     2.9  qed
    2.10  
    2.11  lemma horner_schema:
    2.12 @@ -747,7 +747,7 @@
    2.13      also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then -1 ^ (i div 2) / (real (fact i)) * x ^ i else 0)"
    2.14        unfolding sum_split_even_odd atLeast0LessThan ..
    2.15      also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then -1 ^ (i div 2) / (real (fact i)) else 0) * x ^ i)"
    2.16 -      by (rule setsum_cong2) auto
    2.17 +      by (rule setsum.cong) auto
    2.18      finally show ?thesis by assumption
    2.19    qed } note morph_to_if_power = this
    2.20  
    2.21 @@ -862,7 +862,7 @@
    2.22        also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i)) * x ^ i)"
    2.23          unfolding sum_split_even_odd atLeast0LessThan ..
    2.24        also have "\<dots> = (\<Sum> i = 0 ..< 2 * n. (if even i then 0 else -1 ^ ((i - Suc 0) div 2) / (real (fact i))) * x ^ i)"
    2.25 -        by (rule setsum_cong2) auto
    2.26 +        by (rule setsum.cong) auto
    2.27        finally show ?thesis by assumption
    2.28      qed } note setsum_morph = this
    2.29  
     3.1 --- a/src/HOL/Deriv.thy	Fri Jun 27 22:08:55 2014 +0200
     3.2 +++ b/src/HOL/Deriv.thy	Sat Jun 28 09:16:42 2014 +0200
     3.3 @@ -357,7 +357,7 @@
     3.4      have "((\<lambda>x. f i x * (\<Prod>i\<in>I. f i x)) has_derivative ?P) (at x within s)"
     3.5        using insert by (intro has_derivative_mult) auto
     3.6      also have "?P = (\<lambda>y. \<Sum>i'\<in>insert i I. f' i' y * (\<Prod>j\<in>insert i I - {i'}. f j x))"
     3.7 -      using insert(1,2) by (auto simp add: setsum_right_distrib insert_Diff_if intro!: ext setsum_cong)
     3.8 +      using insert(1,2) by (auto simp add: setsum_right_distrib insert_Diff_if intro!: ext setsum.cong)
     3.9      finally show ?case
    3.10        using insert by simp
    3.11    qed simp  
     4.1 --- a/src/HOL/Fact.thy	Fri Jun 27 22:08:55 2014 +0200
     4.2 +++ b/src/HOL/Fact.thy	Sat Jun 28 09:16:42 2014 +0200
     4.3 @@ -202,7 +202,7 @@
     4.4      also from Suc.hyps have "... = Suc (n + d') * \<Prod>{n + 1..n + d'}" 
     4.5        unfolding div_mult1_eq[of _ "fact (n + d')"] by (simp add: fact_mod)
     4.6      also have "... = \<Prod>{n + 1..n + Suc d'}"
     4.7 -      by (simp add: atLeastAtMostSuc_conv setprod_insert)
     4.8 +      by (simp add: atLeastAtMostSuc_conv setprod.insert)
     4.9      finally show ?case .
    4.10    qed
    4.11    from this `m = n + d` show ?thesis by simp
    4.12 @@ -314,7 +314,7 @@
    4.13    assumes "r \<le> n" shows "fact n div fact (n - r) \<le> n ^ r"
    4.14  proof -
    4.15    have "\<And>r. r \<le> n \<Longrightarrow> \<Prod>{n - r..n} = (n - r) * \<Prod>{Suc (n - r)..n}"
    4.16 -    by (subst setprod_insert[symmetric]) (auto simp: atLeastAtMost_insertL)
    4.17 +    by (subst setprod.insert[symmetric]) (auto simp: atLeastAtMost_insertL)
    4.18    with assms show ?thesis
    4.19      by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono)
    4.20  qed
     5.1 --- a/src/HOL/Groups_Big.thy	Fri Jun 27 22:08:55 2014 +0200
     5.2 +++ b/src/HOL/Groups_Big.thy	Sat Jun 28 09:16:42 2014 +0200
     5.3 @@ -87,6 +87,15 @@
     5.4    shows "F g (A \<union> B) = F g A * F g B"
     5.5    using assms by (simp add: union_inter_neutral)
     5.6  
     5.7 +lemma union_diff2:
     5.8 +  assumes "finite A" and "finite B"
     5.9 +  shows "F g (A \<union> B) = F g (A - B) * F g (B - A) * F g (A \<inter> B)"
    5.10 +proof -
    5.11 +  have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
    5.12 +    by auto
    5.13 +  with assms show ?thesis by simp (subst union_disjoint, auto)+
    5.14 +qed
    5.15 +
    5.16  lemma subset_diff:
    5.17    assumes "B \<subseteq> A" and "finite A"
    5.18    shows "F g A = F g (A - B) * F g B"
    5.19 @@ -139,6 +148,13 @@
    5.20    shows "F (\<lambda>x. g x) A = F (\<lambda>x. h x) B"
    5.21    by (rule cong) (insert assms, simp_all add: simp_implies_def)
    5.22  
    5.23 +lemma reindex_cong:
    5.24 +  assumes "inj_on l B"
    5.25 +  assumes "A = l ` B"
    5.26 +  assumes "\<And>x. x \<in> B \<Longrightarrow> g (l x) = h x"
    5.27 +  shows "F g A = F h B"
    5.28 +  using assms by (simp add: reindex)
    5.29 +
    5.30  lemma UNION_disjoint:
    5.31    assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
    5.32    and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
    5.33 @@ -156,7 +172,7 @@
    5.34  
    5.35  lemma Union_disjoint:
    5.36    assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {}"
    5.37 -  shows "F g (Union C) = F (F g) C"
    5.38 +  shows "F g (Union C) = (F \<circ> F) g C"
    5.39  proof cases
    5.40    assume "finite C"
    5.41    from UNION_disjoint [OF this assms]
    5.42 @@ -257,6 +273,15 @@
    5.43    qed simp
    5.44  qed
    5.45  
    5.46 +lemma reindex_nontrivial:
    5.47 +  assumes "finite A"
    5.48 +  and nz: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<noteq> y \<Longrightarrow> h x = h y \<Longrightarrow> g (h x) = 1"
    5.49 +  shows "F g (h ` A) = F (g \<circ> h) A"
    5.50 +proof (subst reindex_bij_betw_not_neutral [symmetric])
    5.51 +  show "bij_betw h (A - {x \<in> A. (g \<circ> h) x = 1}) (h ` A - h ` {x \<in> A. (g \<circ> h) x = 1})"
    5.52 +    using nz by (auto intro!: inj_onI simp: bij_betw_def)
    5.53 +qed (insert `finite A`, auto)
    5.54 +
    5.55  lemma reindex_bij_witness_not_neutral:
    5.56    assumes fin: "finite S'" "finite T'"
    5.57    assumes witness:
    5.58 @@ -336,6 +361,69 @@
    5.59  apply (cases "B = {}") apply (auto intro: infinite dest: finite_cartesian_productD1)
    5.60  done
    5.61  
    5.62 +lemma inter_restrict:
    5.63 +  assumes "finite A"
    5.64 +  shows "F g (A \<inter> B) = F (\<lambda>x. if x \<in> B then g x else 1) A"
    5.65 +proof -
    5.66 +  let ?g = "\<lambda>x. if x \<in> A \<inter> B then g x else 1"
    5.67 +  have "\<forall>i\<in>A - A \<inter> B. (if i \<in> A \<inter> B then g i else 1) = 1"
    5.68 +   by simp
    5.69 +  moreover have "A \<inter> B \<subseteq> A" by blast
    5.70 +  ultimately have "F ?g (A \<inter> B) = F ?g A" using `finite A`
    5.71 +    by (intro mono_neutral_left) auto
    5.72 +  then show ?thesis by simp
    5.73 +qed
    5.74 +
    5.75 +lemma inter_filter:
    5.76 +  "finite A \<Longrightarrow> F g {x \<in> A. P x} = F (\<lambda>x. if P x then g x else 1) A"
    5.77 +  by (simp add: inter_restrict [symmetric, of A "{x. P x}" g, simplified mem_Collect_eq] Int_def)
    5.78 +
    5.79 +lemma Union_comp:
    5.80 +  assumes "\<forall>A \<in> B. finite A"
    5.81 +    and "\<And>A1 A2 x. A1 \<in> B \<Longrightarrow> A2 \<in> B  \<Longrightarrow> A1 \<noteq> A2 \<Longrightarrow> x \<in> A1 \<Longrightarrow> x \<in> A2 \<Longrightarrow> g x = 1"
    5.82 +  shows "F g (\<Union>B) = (F \<circ> F) g B"
    5.83 +using assms proof (induct B rule: infinite_finite_induct)
    5.84 +  case (infinite A)
    5.85 +  then have "\<not> finite (\<Union>A)" by (blast dest: finite_UnionD)
    5.86 +  with infinite show ?case by simp
    5.87 +next
    5.88 +  case empty then show ?case by simp
    5.89 +next
    5.90 +  case (insert A B)
    5.91 +  then have "finite A" "finite B" "finite (\<Union>B)" "A \<notin> B"
    5.92 +    and "\<forall>x\<in>A \<inter> \<Union>B. g x = 1"
    5.93 +    and H: "F g (\<Union>B) = (F o F) g B" by auto
    5.94 +  then have "F g (A \<union> \<Union>B) = F g A * F g (\<Union>B)"
    5.95 +    by (simp add: union_inter_neutral)
    5.96 +  with `finite B` `A \<notin> B` show ?case
    5.97 +    by (simp add: H)
    5.98 +qed
    5.99 +
   5.100 +lemma commute:
   5.101 +  "F (\<lambda>i. F (g i) B) A = F (\<lambda>j. F (\<lambda>i. g i j) A) B"
   5.102 +  unfolding cartesian_product
   5.103 +  by (rule reindex_bij_witness [where i = "\<lambda>(i, j). (j, i)" and j = "\<lambda>(i, j). (j, i)"]) auto
   5.104 +
   5.105 +lemma commute_restrict:
   5.106 +  "finite A \<Longrightarrow> finite B \<Longrightarrow>
   5.107 +    F (\<lambda>x. F (g x) {y. y \<in> B \<and> R x y}) A = F (\<lambda>y. F (\<lambda>x. g x y) {x. x \<in> A \<and> R x y}) B"
   5.108 +  by (simp add: inter_filter) (rule commute)
   5.109 +
   5.110 +lemma Plus:
   5.111 +  fixes A :: "'b set" and B :: "'c set"
   5.112 +  assumes fin: "finite A" "finite B"
   5.113 +  shows "F g (A <+> B) = F (g \<circ> Inl) A * F (g \<circ> Inr) B"
   5.114 +proof -
   5.115 +  have "A <+> B = Inl ` A \<union> Inr ` B" by auto
   5.116 +  moreover from fin have "finite (Inl ` A :: ('b + 'c) set)" "finite (Inr ` B :: ('b + 'c) set)"
   5.117 +    by auto
   5.118 +  moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('b + 'c) set)" by auto
   5.119 +  moreover have "inj_on (Inl :: 'b \<Rightarrow> 'b + 'c) A" "inj_on (Inr :: 'c \<Rightarrow> 'b + 'c) B"
   5.120 +    by (auto intro: inj_onI)
   5.121 +  ultimately show ?thesis using fin
   5.122 +    by (simp add: union_disjoint reindex)
   5.123 +qed
   5.124 +
   5.125  end
   5.126  
   5.127  notation times (infixl "*" 70)
   5.128 @@ -410,191 +498,27 @@
   5.129  in [(@{const_syntax setsum}, K setsum_tr')] end
   5.130  *}
   5.131  
   5.132 -text {* TODO These are candidates for generalization *}
   5.133 +text {* TODO generalization candidates *}
   5.134  
   5.135 -context comm_monoid_add
   5.136 -begin
   5.137 -
   5.138 -lemma setsum_reindex_id: 
   5.139 -  "inj_on f B \<Longrightarrow> setsum f B = setsum id (f ` B)"
   5.140 -  by (simp add: setsum.reindex)
   5.141 -
   5.142 -lemma setsum_reindex_nonzero:
   5.143 +lemma setsum_image_gen:
   5.144    assumes fS: "finite S"
   5.145 -  and nz: "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> x \<noteq> y \<Longrightarrow> f x = f y \<Longrightarrow> h (f x) = 0"
   5.146 -  shows "setsum h (f ` S) = setsum (h \<circ> f) S"
   5.147 -proof (subst setsum.reindex_bij_betw_not_neutral[symmetric])
   5.148 -  show "bij_betw f (S - {x\<in>S. h (f x) = 0}) (f`S - f`{x\<in>S. h (f x) = 0})"
   5.149 -    using nz by (auto intro!: inj_onI simp: bij_betw_def)
   5.150 -qed (insert fS, auto)
   5.151 -
   5.152 -lemma setsum_cong2:
   5.153 -  "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
   5.154 -  by (auto intro: setsum.cong)
   5.155 -
   5.156 -lemma setsum_reindex_cong:
   5.157 -   "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|] 
   5.158 -    ==> setsum h B = setsum g A"
   5.159 -  by (simp add: setsum.reindex)
   5.160 -
   5.161 -lemma setsum_restrict_set:
   5.162 -  assumes fA: "finite A"
   5.163 -  shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
   5.164 +  shows "setsum g S = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
   5.165  proof-
   5.166 -  from fA have fab: "finite (A \<inter> B)" by auto
   5.167 -  have aba: "A \<inter> B \<subseteq> A" by blast
   5.168 -  let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
   5.169 -  from setsum.mono_neutral_left [OF fA aba, of ?g]
   5.170 -  show ?thesis by simp
   5.171 +  { fix x assume "x \<in> S" then have "{y. y\<in> f`S \<and> f x = y} = {f x}" by auto }
   5.172 +  hence "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
   5.173 +    by simp
   5.174 +  also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
   5.175 +    by (rule setsum.commute_restrict [OF fS finite_imageI [OF fS]])
   5.176 +  finally show ?thesis .
   5.177  qed
   5.178  
   5.179 -lemma setsum_Union_disjoint:
   5.180 -  assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}"
   5.181 -  shows "setsum f (Union C) = setsum (setsum f) C"
   5.182 -  using assms by (fact setsum.Union_disjoint)
   5.183 -
   5.184 -lemma setsum_cartesian_product:
   5.185 -  "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
   5.186 -  by (fact setsum.cartesian_product)
   5.187 -
   5.188 -lemma setsum_UNION_zero:
   5.189 -  assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
   5.190 -  and f0: "\<And>T1 T2 x. T1\<in>S \<Longrightarrow> T2\<in>S \<Longrightarrow> T1 \<noteq> T2 \<Longrightarrow> x \<in> T1 \<Longrightarrow> x \<in> T2 \<Longrightarrow> f x = 0"
   5.191 -  shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
   5.192 -  using fSS f0
   5.193 -proof(induct rule: finite_induct[OF fS])
   5.194 -  case 1 thus ?case by simp
   5.195 -next
   5.196 -  case (2 T F)
   5.197 -  then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F" 
   5.198 -    and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
   5.199 -  from fTF have fUF: "finite (\<Union>F)" by auto
   5.200 -  from "2.prems" TF fTF
   5.201 -  show ?case 
   5.202 -    by (auto simp add: H [symmetric] intro: setsum.union_inter_neutral [OF fTF(1) fUF, of f])
   5.203 -qed
   5.204 -
   5.205 -text {* Commuting outer and inner summation *}
   5.206 -
   5.207 -lemma setsum_commute:
   5.208 -  "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
   5.209 -  unfolding setsum_cartesian_product
   5.210 -  by (rule setsum.reindex_bij_witness[where i="\<lambda>(i, j). (j, i)" and j="\<lambda>(i, j). (j, i)"]) auto
   5.211 -
   5.212 -lemma setsum_Plus:
   5.213 -  fixes A :: "'a set" and B :: "'b set"
   5.214 -  assumes fin: "finite A" "finite B"
   5.215 -  shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
   5.216 -proof -
   5.217 -  have "A <+> B = Inl ` A \<union> Inr ` B" by auto
   5.218 -  moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
   5.219 -    by auto
   5.220 -  moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
   5.221 -  moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
   5.222 -  ultimately show ?thesis using fin by(simp add: setsum.union_disjoint setsum.reindex)
   5.223 -qed
   5.224 -
   5.225 -end
   5.226 -
   5.227 -text {* TODO These are legacy *}
   5.228 -
   5.229 -lemma setsum_empty:
   5.230 -  "setsum f {} = 0"
   5.231 -  by (fact setsum.empty)
   5.232 -
   5.233 -lemma setsum_insert:
   5.234 -  "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
   5.235 -  by (fact setsum.insert)
   5.236 -
   5.237 -lemma setsum_infinite:
   5.238 -  "~ finite A ==> setsum f A = 0"
   5.239 -  by (fact setsum.infinite)
   5.240 -
   5.241 -lemma setsum_reindex:
   5.242 -  "inj_on f B \<Longrightarrow> setsum h (f ` B) = setsum (h \<circ> f) B"
   5.243 -  by (fact setsum.reindex)
   5.244 -
   5.245 -lemma setsum_cong:
   5.246 -  "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
   5.247 -  by (fact setsum.cong)
   5.248 -
   5.249 -lemma strong_setsum_cong:
   5.250 -  "A = B ==> (!!x. x:B =simp=> f x = g x)
   5.251 -   ==> setsum (%x. f x) A = setsum (%x. g x) B"
   5.252 -  by (fact setsum.strong_cong)
   5.253 -
   5.254 -lemmas setsum_0 = setsum.neutral_const
   5.255 -lemmas setsum_0' = setsum.neutral
   5.256 -
   5.257 -lemma setsum_Un_Int: "finite A ==> finite B ==>
   5.258 -  setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
   5.259 -  -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
   5.260 -  by (fact setsum.union_inter)
   5.261 -
   5.262 -lemma setsum_Un_disjoint: "finite A ==> finite B
   5.263 -  ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
   5.264 -  by (fact setsum.union_disjoint)
   5.265 -
   5.266 -lemma setsum_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>
   5.267 -    setsum f A = setsum f (A - B) + setsum f B"
   5.268 -  by (fact setsum.subset_diff)
   5.269 -
   5.270 -lemma setsum_mono_zero_left: 
   5.271 -  "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 0 \<rbrakk> \<Longrightarrow> setsum f S = setsum f T"
   5.272 -  by (fact setsum.mono_neutral_left)
   5.273 -
   5.274 -lemmas setsum_mono_zero_right = setsum.mono_neutral_right
   5.275 -
   5.276 -lemma setsum_mono_zero_cong_left: 
   5.277 -  "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 0; \<And>x. x \<in> S \<Longrightarrow> f x = g x \<rbrakk>
   5.278 -  \<Longrightarrow> setsum f S = setsum g T"
   5.279 -  by (fact setsum.mono_neutral_cong_left)
   5.280 -
   5.281 -lemmas setsum_mono_zero_cong_right = setsum.mono_neutral_cong_right
   5.282 -
   5.283 -lemma setsum_delta: "finite S \<Longrightarrow>
   5.284 -  setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
   5.285 -  by (fact setsum.delta)
   5.286 -
   5.287 -lemma setsum_delta': "finite S \<Longrightarrow>
   5.288 -  setsum (\<lambda>k. if a = k then b k else 0) S = (if a\<in> S then b a else 0)"
   5.289 -  by (fact setsum.delta')
   5.290 -
   5.291 -lemma setsum_cases:
   5.292 -  assumes "finite A"
   5.293 -  shows "setsum (\<lambda>x. if P x then f x else g x) A =
   5.294 -         setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
   5.295 -  using assms by (fact setsum.If_cases)
   5.296 -
   5.297 -(*But we can't get rid of finite I. If infinite, although the rhs is 0, 
   5.298 -  the lhs need not be, since UNION I A could still be finite.*)
   5.299 -lemma setsum_UN_disjoint:
   5.300 -  assumes "finite I" and "ALL i:I. finite (A i)"
   5.301 -    and "ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}"
   5.302 -  shows "setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
   5.303 -  using assms by (fact setsum.UNION_disjoint)
   5.304 -
   5.305 -(*But we can't get rid of finite A. If infinite, although the lhs is 0, 
   5.306 -  the rhs need not be, since SIGMA A B could still be finite.*)
   5.307 -lemma setsum_Sigma:
   5.308 -  assumes "finite A" and  "ALL x:A. finite (B x)"
   5.309 -  shows "(\<Sum>x\<in>A. (\<Sum>y\<in>B x. f x y)) = (\<Sum>(x,y)\<in>(SIGMA x:A. B x). f x y)"
   5.310 -  using assms by (fact setsum.Sigma)
   5.311 -
   5.312 -lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
   5.313 -  by (fact setsum.distrib)
   5.314 -
   5.315 -lemma setsum_Un_zero:  
   5.316 -  "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 0 \<rbrakk> \<Longrightarrow>
   5.317 -  setsum f (S \<union> T) = setsum f S + setsum f T"
   5.318 -  by (fact setsum.union_inter_neutral)
   5.319  
   5.320  subsubsection {* Properties in more restricted classes of structures *}
   5.321  
   5.322  lemma setsum_Un: "finite A ==> finite B ==>
   5.323    (setsum f (A Un B) :: 'a :: ab_group_add) =
   5.324     setsum f A + setsum f B - setsum f (A Int B)"
   5.325 -by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
   5.326 +by (subst setsum.union_inter [symmetric], auto simp add: algebra_simps)
   5.327  
   5.328  lemma setsum_Un2:
   5.329    assumes "finite (A \<union> B)"
   5.330 @@ -602,7 +526,7 @@
   5.331  proof -
   5.332    have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
   5.333      by auto
   5.334 -  with assms show ?thesis by simp (subst setsum_Un_disjoint, auto)+
   5.335 +  with assms show ?thesis by simp (subst setsum.union_disjoint, auto)+
   5.336  qed
   5.337  
   5.338  lemma setsum_diff1: "finite A \<Longrightarrow>
   5.339 @@ -664,12 +588,12 @@
   5.340    have "setsum f A = setsum f ((A-{a}) \<union> {a})"
   5.341      by(simp add:insert_absorb[OF `a:A`])
   5.342    also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
   5.343 -    using `finite A` by(subst setsum_Un_disjoint) auto
   5.344 +    using `finite A` by(subst setsum.union_disjoint) auto
   5.345    also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
   5.346      by(rule setsum_mono)(simp add: assms(2))
   5.347    also have "setsum f {a} < setsum g {a}" using a by simp
   5.348    also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
   5.349 -    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
   5.350 +    using `finite A` by(subst setsum.union_disjoint[symmetric]) auto
   5.351    also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
   5.352    finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono)
   5.353  qed
   5.354 @@ -685,7 +609,7 @@
   5.355  lemma setsum_subtractf:
   5.356    "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
   5.357      setsum f A - setsum g A"
   5.358 -  using setsum_addf [of f "- g" A] by (simp add: setsum_negf)
   5.359 +  using setsum.distrib [of f "- g" A] by (simp add: setsum_negf)
   5.360  
   5.361  lemma setsum_nonneg:
   5.362    assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
   5.363 @@ -747,11 +671,32 @@
   5.364    have "setsum f A \<le> setsum f A + setsum f (B-A)"
   5.365      by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
   5.366    also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
   5.367 -    by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
   5.368 +    by (simp add: setsum.union_disjoint del:Un_Diff_cancel)
   5.369    also have "A \<union> (B-A) = B" using sub by blast
   5.370    finally show ?thesis .
   5.371  qed
   5.372  
   5.373 +lemma setsum_le_included:
   5.374 +  fixes f :: "'a \<Rightarrow> 'b::ordered_comm_monoid_add"
   5.375 +  assumes "finite s" "finite t"
   5.376 +  and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"
   5.377 +  shows "setsum f s \<le> setsum g t"
   5.378 +proof -
   5.379 +  have "setsum f s \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) s"
   5.380 +  proof (rule setsum_mono)
   5.381 +    fix y assume "y \<in> s"
   5.382 +    with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto
   5.383 +    with assms show "f y \<le> setsum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
   5.384 +      using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro]
   5.385 +      by (auto intro!: setsum_mono2)
   5.386 +  qed
   5.387 +  also have "... \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) (i ` t)"
   5.388 +    using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg)
   5.389 +  also have "... \<le> setsum g t"
   5.390 +    using assms by (auto simp: setsum_image_gen[symmetric])
   5.391 +  finally show ?thesis .
   5.392 +qed
   5.393 +
   5.394  lemma setsum_mono3: "finite B ==> A <= B ==> 
   5.395      ALL x: B - A. 
   5.396        0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
   5.397 @@ -762,7 +707,7 @@
   5.398    apply simp
   5.399    apply (rule add_left_mono)
   5.400    apply (erule setsum_nonneg)
   5.401 -  apply (subst setsum_Un_disjoint [THEN sym])
   5.402 +  apply (subst setsum.union_disjoint [THEN sym])
   5.403    apply (erule finite_subset, assumption)
   5.404    apply (rule finite_subset)
   5.405    prefer 2
   5.406 @@ -865,27 +810,21 @@
   5.407    case False thus ?thesis by simp
   5.408  qed
   5.409  
   5.410 -lemma setsum_diff1'[rule_format]:
   5.411 -  "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
   5.412 -apply (erule finite_induct[where F=A and P="% A. (a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x))"])
   5.413 -apply (auto simp add: insert_Diff_if add_ac)
   5.414 -done
   5.415 -
   5.416  lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
   5.417    shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
   5.418 -unfolding setsum_diff1'[OF assms] by auto
   5.419 +  unfolding setsum.remove [OF assms] by auto
   5.420  
   5.421  lemma setsum_product:
   5.422    fixes f :: "'a => ('b::semiring_0)"
   5.423    shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
   5.424 -  by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
   5.425 +  by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum.commute)
   5.426  
   5.427  lemma setsum_mult_setsum_if_inj:
   5.428  fixes f :: "'a => ('b::semiring_0)"
   5.429  shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
   5.430    setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
   5.431 -by(auto simp: setsum_product setsum_cartesian_product
   5.432 -        intro!:  setsum_reindex_cong[symmetric])
   5.433 +by(auto simp: setsum_product setsum.cartesian_product
   5.434 +        intro!:  setsum.reindex_cong[symmetric])
   5.435  
   5.436  lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
   5.437  apply (case_tac "finite A")
   5.438 @@ -909,7 +848,7 @@
   5.439  lemma setsum_Un_nat: "finite A ==> finite B ==>
   5.440    (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
   5.441    -- {* For the natural numbers, we have subtraction. *}
   5.442 -by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
   5.443 +by (subst setsum.union_inter [symmetric], auto simp add: algebra_simps)
   5.444  
   5.445  lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
   5.446    (if a:A then setsum f A - f a else setsum f A)"
   5.447 @@ -993,7 +932,7 @@
   5.448    shows "card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
   5.449  proof -
   5.450    have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)" by simp
   5.451 -  with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant)
   5.452 +  with assms show ?thesis by (simp add: card_eq_setsum setsum.UNION_disjoint del: setsum_constant)
   5.453  qed
   5.454  
   5.455  lemma card_Union_disjoint:
   5.456 @@ -1004,13 +943,32 @@
   5.457  apply simp_all
   5.458  done
   5.459  
   5.460 +lemma setsum_multicount_gen:
   5.461 +  assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
   5.462 +  shows "setsum (\<lambda>i. (card {j\<in>t. R i j})) s = setsum k t" (is "?l = ?r")
   5.463 +proof-
   5.464 +  have "?l = setsum (\<lambda>i. setsum (\<lambda>x.1) {j\<in>t. R i j}) s" by auto
   5.465 +  also have "\<dots> = ?r" unfolding setsum.commute_restrict [OF assms(1-2)]
   5.466 +    using assms(3) by auto
   5.467 +  finally show ?thesis .
   5.468 +qed
   5.469 +
   5.470 +lemma setsum_multicount:
   5.471 +  assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
   5.472 +  shows "setsum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
   5.473 +proof-
   5.474 +  have "?l = setsum (\<lambda>i. k) T" by (rule setsum_multicount_gen) (auto simp: assms)
   5.475 +  also have "\<dots> = ?r" by (simp add: mult_commute)
   5.476 +  finally show ?thesis by auto
   5.477 +qed
   5.478 +
   5.479  
   5.480  subsubsection {* Cardinality of products *}
   5.481  
   5.482  lemma card_SigmaI [simp]:
   5.483    "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
   5.484    \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
   5.485 -by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant)
   5.486 +by(simp add: card_eq_setsum setsum.Sigma del:setsum_constant)
   5.487  
   5.488  (*
   5.489  lemma SigmaI_insert: "y \<notin> A ==>
   5.490 @@ -1075,124 +1033,6 @@
   5.491    "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
   5.492    "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
   5.493  
   5.494 -text {* TODO These are candidates for generalization *}
   5.495 -
   5.496 -context comm_monoid_mult
   5.497 -begin
   5.498 -
   5.499 -lemma setprod_reindex_id:
   5.500 -  "inj_on f B ==> setprod f B = setprod id (f ` B)"
   5.501 -  by (auto simp add: setprod.reindex)
   5.502 -
   5.503 -lemma setprod_reindex_cong:
   5.504 -  "inj_on f A ==> B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
   5.505 -  by (frule setprod.reindex, simp)
   5.506 -
   5.507 -lemma strong_setprod_reindex_cong:
   5.508 -  "inj_on f A \<Longrightarrow> B = f ` A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x) \<Longrightarrow> setprod h B = setprod g A"
   5.509 -  by (subst setprod.reindex_bij_betw[symmetric, where h=f])
   5.510 -     (auto simp: bij_betw_def)
   5.511 -
   5.512 -lemma setprod_Union_disjoint:
   5.513 -  assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}" 
   5.514 -  shows "setprod f (Union C) = setprod (setprod f) C"
   5.515 -  using assms by (fact setprod.Union_disjoint)
   5.516 -
   5.517 -text{*Here we can eliminate the finiteness assumptions, by cases.*}
   5.518 -lemma setprod_cartesian_product:
   5.519 -  "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
   5.520 -  by (fact setprod.cartesian_product)
   5.521 -
   5.522 -lemma setprod_Un2:
   5.523 -  assumes "finite (A \<union> B)"
   5.524 -  shows "setprod f (A \<union> B) = setprod f (A - B) * setprod f (B - A) * setprod f (A \<inter> B)"
   5.525 -proof -
   5.526 -  have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
   5.527 -    by auto
   5.528 -  with assms show ?thesis by simp (subst setprod.union_disjoint, auto)+
   5.529 -qed
   5.530 -
   5.531 -end
   5.532 -
   5.533 -text {* TODO These are legacy *}
   5.534 -
   5.535 -lemma setprod_empty: "setprod f {} = 1"
   5.536 -  by (fact setprod.empty)
   5.537 -
   5.538 -lemma setprod_insert: "[| finite A; a \<notin> A |] ==>
   5.539 -    setprod f (insert a A) = f a * setprod f A"
   5.540 -  by (fact setprod.insert)
   5.541 -
   5.542 -lemma setprod_infinite: "~ finite A ==> setprod f A = 1"
   5.543 -  by (fact setprod.infinite)
   5.544 -
   5.545 -lemma setprod_reindex:
   5.546 -  "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
   5.547 -  by (fact setprod.reindex)
   5.548 -
   5.549 -lemma setprod_cong:
   5.550 -  "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
   5.551 -  by (fact setprod.cong)
   5.552 -
   5.553 -lemma strong_setprod_cong:
   5.554 -  "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
   5.555 -  by (fact setprod.strong_cong)
   5.556 -
   5.557 -lemma setprod_Un_one:
   5.558 -  "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 1 \<rbrakk>
   5.559 -  \<Longrightarrow> setprod f (S \<union> T) = setprod f S  * setprod f T"
   5.560 -  by (fact setprod.union_inter_neutral)
   5.561 -
   5.562 -lemmas setprod_1 = setprod.neutral_const
   5.563 -lemmas setprod_1' = setprod.neutral
   5.564 -
   5.565 -lemma setprod_Un_Int: "finite A ==> finite B
   5.566 -    ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
   5.567 -  by (fact setprod.union_inter)
   5.568 -
   5.569 -lemma setprod_Un_disjoint: "finite A ==> finite B
   5.570 -  ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
   5.571 -  by (fact setprod.union_disjoint)
   5.572 -
   5.573 -lemma setprod_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>
   5.574 -    setprod f A = setprod f (A - B) * setprod f B"
   5.575 -  by (fact setprod.subset_diff)
   5.576 -
   5.577 -lemma setprod_mono_one_left:
   5.578 -  "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 1 \<rbrakk> \<Longrightarrow> setprod f S = setprod f T"
   5.579 -  by (fact setprod.mono_neutral_left)
   5.580 -
   5.581 -lemmas setprod_mono_one_right = setprod.mono_neutral_right
   5.582 -
   5.583 -lemma setprod_mono_one_cong_left: 
   5.584 -  "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1; \<And>x. x \<in> S \<Longrightarrow> f x = g x \<rbrakk>
   5.585 -  \<Longrightarrow> setprod f S = setprod g T"
   5.586 -  by (fact setprod.mono_neutral_cong_left)
   5.587 -
   5.588 -lemmas setprod_mono_one_cong_right = setprod.mono_neutral_cong_right
   5.589 -
   5.590 -lemma setprod_delta: "finite S \<Longrightarrow>
   5.591 -  setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
   5.592 -  by (fact setprod.delta)
   5.593 -
   5.594 -lemma setprod_delta': "finite S \<Longrightarrow>
   5.595 -  setprod (\<lambda>k. if a = k then b k else 1) S = (if a\<in> S then b a else 1)"
   5.596 -  by (fact setprod.delta')
   5.597 -
   5.598 -lemma setprod_UN_disjoint:
   5.599 -    "finite I ==> (ALL i:I. finite (A i)) ==>
   5.600 -        (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
   5.601 -      setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
   5.602 -  by (fact setprod.UNION_disjoint)
   5.603 -
   5.604 -lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
   5.605 -    (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
   5.606 -    (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
   5.607 -  by (fact setprod.Sigma)
   5.608 -
   5.609 -lemma setprod_timesf: "setprod (\<lambda>x. f x * g x) A = setprod f A * setprod g A"
   5.610 -  by (fact setprod.distrib)
   5.611 -
   5.612  
   5.613  subsubsection {* Properties in more restricted classes of structures *}
   5.614  
   5.615 @@ -1210,7 +1050,7 @@
   5.616  lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
   5.617    (setprod f (A Un B) :: 'a ::{field})
   5.618     = setprod f A * setprod f B / setprod f (A Int B)"
   5.619 -by (subst setprod_Un_Int [symmetric], auto)
   5.620 +by (subst setprod.union_inter [symmetric], auto)
   5.621  
   5.622  lemma setprod_nonneg [rule_format]:
   5.623     "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
   5.624 @@ -1238,9 +1078,9 @@
   5.625           "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
   5.626  apply (erule ssubst)
   5.627  apply (subst divide_inverse)
   5.628 -apply (subst setprod_timesf)
   5.629 +apply (subst setprod.distrib)
   5.630  apply (subst setprod_inversef, assumption+, rule refl)
   5.631 -apply (rule setprod_cong, rule refl)
   5.632 +apply (rule setprod.cong, rule refl)
   5.633  apply (subst divide_inverse, auto)
   5.634  done
   5.635  
   5.636 @@ -1257,8 +1097,8 @@
   5.637    "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
   5.638    apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
   5.639    apply (unfold dvd_def, blast)
   5.640 -  apply (subst setprod_Un_disjoint [symmetric])
   5.641 -  apply (auto elim: finite_subset intro: setprod_cong)
   5.642 +  apply (subst setprod.union_disjoint [symmetric])
   5.643 +  apply (auto elim: finite_subset intro: setprod.cong)
   5.644  done
   5.645  
   5.646  lemma setprod_dvd_setprod_subset2:
   5.647 @@ -1290,7 +1130,7 @@
   5.648    proof (induct A rule: finite_subset_induct)
   5.649      case (insert a F)
   5.650      thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
   5.651 -      unfolding setprod_insert[OF insert(1,3)]
   5.652 +      unfolding setprod.insert[OF insert(1,3)]
   5.653        using assms[rule_format,OF insert(2)] insert
   5.654        by (auto intro: mult_mono)
   5.655    qed auto
     6.1 --- a/src/HOL/Hoare_Parallel/OG_Examples.thy	Fri Jun 27 22:08:55 2014 +0200
     6.2 +++ b/src/HOL/Hoare_Parallel/OG_Examples.thy	Sat Jun 28 09:16:42 2014 +0200
     6.3 @@ -534,9 +534,9 @@
     6.4   COEND 
     6.5   \<lbrace>\<acute>x=n\<rbrace>"
     6.6  apply oghoare
     6.7 -apply (simp_all cong del: strong_setsum_cong)
     6.8 +apply (simp_all cong del: setsum.strong_cong)
     6.9  apply (tactic {* ALLGOALS (clarify_tac @{context}) *})
    6.10 -apply (simp_all cong del: strong_setsum_cong)
    6.11 +apply (simp_all cong del: setsum.strong_cong)
    6.12     apply(erule (1) Example2_lemma2)
    6.13    apply(erule (1) Example2_lemma2)
    6.14   apply(erule (1) Example2_lemma2)
     7.1 --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Fri Jun 27 22:08:55 2014 +0200
     7.2 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int1_ITP.thy	Sat Jun 28 09:16:42 2014 +0200
     7.3 @@ -182,12 +182,12 @@
     7.4    have "setsum f A = setsum f ((A-{a}) \<union> {a})"
     7.5      by(simp add:insert_absorb[OF `a:A`])
     7.6    also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
     7.7 -    using `finite A` by(subst setsum_Un_disjoint) auto
     7.8 +    using `finite A` by(subst setsum.union_disjoint) auto
     7.9    also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
    7.10      by(rule setsum_mono)(simp add: assms(2))
    7.11    also have "setsum f {a} < setsum g {a}" using a by simp
    7.12    also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
    7.13 -    using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
    7.14 +    using `finite A` by(subst setsum.union_disjoint[symmetric]) auto
    7.15    also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
    7.16    finally show ?thesis by (metis add_right_mono add_strict_left_mono)
    7.17  qed
    7.18 @@ -241,7 +241,7 @@
    7.19        also have "\<dots> \<le> (\<Sum>y\<in>?X'-{u}. m(?f y)+1)"
    7.20          by(simp add: setsum_mono3[OF _ `?Y'\<inter>?X-{u} <= ?X'-{u}`])
    7.21        also have "\<dots> + (\<Sum>y\<in>{u}. m(?f y)+1)= (\<Sum>y\<in>(?X'-{u}) \<union> {u}. m(?f y)+1)"
    7.22 -        using `u:?X'` by(subst setsum_Un_disjoint[symmetric]) auto
    7.23 +        using `u:?X'` by(subst setsum.union_disjoint[symmetric]) auto
    7.24        also have "\<dots> = (\<Sum>x\<in>?X'. m(?f x)+1)"
    7.25          using `u : ?X'` by(simp add:insert_absorb)
    7.26        finally show ?thesis by (blast intro: add_right_mono)
     8.1 --- a/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy	Fri Jun 27 22:08:55 2014 +0200
     8.2 +++ b/src/HOL/IMP/Abs_Int_ITP/Abs_Int3_ITP.thy	Sat Jun 28 09:16:42 2014 +0200
     8.3 @@ -319,7 +319,7 @@
     8.4    have "(\<Sum>y\<in>?Y. m_ivl(?g y)) = (\<Sum>y\<in>(?Y-?X) \<union> (?Y\<inter>?X). m_ivl(?g y))"
     8.5      by (metis Un_Diff_Int)
     8.6    also have "\<dots> = (\<Sum>y\<in>?Y-?X. m_ivl(?g y)) + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))"
     8.7 -    by(subst setsum_Un_disjoint) auto
     8.8 +    by(subst setsum.union_disjoint) auto
     8.9    also have "(\<Sum>y\<in>?Y-?X. m_ivl(?g y)) = 0" using 0 by simp
    8.10    also have "0 + (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y)) = (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?g y))" by simp
    8.11    also have "\<dots> \<le> (\<Sum>y\<in>?Y\<inter>?X. m_ivl(?f y))"
     9.1 --- a/src/HOL/Library/BigO.thy	Fri Jun 27 22:08:55 2014 +0200
     9.2 +++ b/src/HOL/Library/BigO.thy	Sat Jun 28 09:16:42 2014 +0200
     9.3 @@ -641,7 +641,8 @@
     9.4    apply (erule ssubst)
     9.5    apply (erule bigo_setsum3)
     9.6    apply (rule ext)
     9.7 -  apply (rule setsum_cong2)
     9.8 +  apply (rule setsum.cong)
     9.9 +  apply (rule refl)
    9.10    apply (subst abs_of_nonneg)
    9.11    apply auto
    9.12    done
    10.1 --- a/src/HOL/Library/Convex.thy	Fri Jun 27 22:08:55 2014 +0200
    10.2 +++ b/src/HOL/Library/Convex.thy	Sat Jun 28 09:16:42 2014 +0200
    10.3 @@ -200,7 +200,7 @@
    10.4      then have card: "card ({1 :: nat .. 2} \<inter> - {x. x = 1}) = 1"
    10.5        by simp
    10.6      then have "setsum ?u {1 .. 2} = 1"
    10.7 -      using setsum_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
    10.8 +      using setsum.If_cases[of "{(1 :: nat) .. 2}" "\<lambda> x. x = 1" "\<lambda> x. \<mu>" "\<lambda> x. 1 - \<mu>"]
    10.9        by auto
   10.10      with asm[rule_format, of "2" ?u ?x] have s: "(\<Sum>j \<in> {1..2}. ?u j *\<^sub>R ?x j) \<in> s"
   10.11        using mu xy by auto
   10.12 @@ -270,7 +270,7 @@
   10.13      by simp
   10.14    show "(\<Sum>x\<in>t. u x *\<^sub>R x) \<in> s"
   10.15     using sum[THEN spec[where x="\<lambda>x. if x\<in>t then u x else 0"]] as *
   10.16 -   by (auto simp: assms setsum_cases if_distrib if_distrib_arg)
   10.17 +   by (auto simp: assms setsum.If_cases if_distrib if_distrib_arg)
   10.18  qed (erule_tac x=s in allE, erule_tac x=u in allE, auto)
   10.19  
   10.20  
    11.1 --- a/src/HOL/Library/FSet.thy	Fri Jun 27 22:08:55 2014 +0200
    11.2 +++ b/src/HOL/Library/FSet.thy	Sat Jun 28 09:16:42 2014 +0200
    11.3 @@ -1002,7 +1002,7 @@
    11.4  
    11.5  lemma fset_size_o_map: "inj f \<Longrightarrow> size_fset g \<circ> fimage f = size_fset (g \<circ> f)"
    11.6    unfolding size_fset_def fimage_def
    11.7 -  by (auto simp: Abs_fset_inverse setsum_reindex_cong[OF subset_inj_on[OF _ top_greatest]])
    11.8 +  by (auto simp: Abs_fset_inverse setsum.reindex_cong[OF subset_inj_on[OF _ top_greatest]])
    11.9  
   11.10  setup {*
   11.11  BNF_LFP_Size.register_size_global @{type_name fset} @{const_name size_fset}
    12.1 --- a/src/HOL/Library/Formal_Power_Series.thy	Fri Jun 27 22:08:55 2014 +0200
    12.2 +++ b/src/HOL/Library/Formal_Power_Series.thy	Sat Jun 28 09:16:42 2014 +0200
    12.3 @@ -143,7 +143,7 @@
    12.4      and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
    12.5    shows "(\<Sum>j=0..k. \<Sum>i=0..j. f i (j - i) (n - j)) =
    12.6           (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
    12.7 -  by (induct k) (simp_all add: Suc_diff_le setsum_addf add_assoc)
    12.8 +  by (induct k) (simp_all add: Suc_diff_le setsum.distrib add_assoc)
    12.9  
   12.10  instance fps :: (semiring_0) semigroup_mult
   12.11  proof
   12.12 @@ -194,8 +194,8 @@
   12.13  instance fps :: (semiring_1) monoid_mult
   12.14  proof
   12.15    fix a :: "'a fps"
   12.16 -  show "1 * a = a" by (simp add: fps_ext fps_mult_nth mult_delta_left setsum_delta)
   12.17 -  show "a * 1 = a" by (simp add: fps_ext fps_mult_nth mult_delta_right setsum_delta')
   12.18 +  show "1 * a = a" by (simp add: fps_ext fps_mult_nth mult_delta_left setsum.delta)
   12.19 +  show "a * 1 = a" by (simp add: fps_ext fps_mult_nth mult_delta_right setsum.delta')
   12.20  qed
   12.21  
   12.22  instance fps :: (cancel_semigroup_add) cancel_semigroup_add
   12.23 @@ -235,9 +235,9 @@
   12.24  proof
   12.25    fix a b c :: "'a fps"
   12.26    show "(a + b) * c = a * c + b * c"
   12.27 -    by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum_addf)
   12.28 +    by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum.distrib)
   12.29    show "a * (b + c) = a * b + a * c"
   12.30 -    by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum_addf)
   12.31 +    by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum.distrib)
   12.32  qed
   12.33  
   12.34  instance fps :: (semiring_0) semiring_0
   12.35 @@ -306,7 +306,7 @@
   12.36    by (simp add: fps_ext)
   12.37  
   12.38  lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
   12.39 -  by (simp add: fps_eq_iff fps_mult_nth setsum_0')
   12.40 +  by (simp add: fps_eq_iff fps_mult_nth setsum.neutral)
   12.41  
   12.42  lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
   12.43      Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
   12.44 @@ -318,17 +318,17 @@
   12.45  
   12.46  lemma fps_const_mult_left: "fps_const (c::'a::semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
   12.47    unfolding fps_eq_iff fps_mult_nth
   12.48 -  by (simp add: fps_const_def mult_delta_left setsum_delta)
   12.49 +  by (simp add: fps_const_def mult_delta_left setsum.delta)
   12.50  
   12.51  lemma fps_const_mult_right: "f * fps_const (c::'a::semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
   12.52    unfolding fps_eq_iff fps_mult_nth
   12.53 -  by (simp add: fps_const_def mult_delta_right setsum_delta')
   12.54 +  by (simp add: fps_const_def mult_delta_right setsum.delta')
   12.55  
   12.56  lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
   12.57 -  by (simp add: fps_mult_nth mult_delta_left setsum_delta)
   12.58 +  by (simp add: fps_mult_nth mult_delta_left setsum.delta)
   12.59  
   12.60  lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
   12.61 -  by (simp add: fps_mult_nth mult_delta_right setsum_delta')
   12.62 +  by (simp add: fps_mult_nth mult_delta_right setsum.delta')
   12.63  
   12.64  subsection {* Formal power series form an integral domain*}
   12.65  
   12.66 @@ -350,9 +350,9 @@
   12.67    have "(a * b) $ (i+j) = (\<Sum>k=0..i+j. a$k * b$(i+j-k))"
   12.68      by (rule fps_mult_nth)
   12.69    also have "\<dots> = (a$i * b$(i+j-i)) + (\<Sum>k\<in>{0..i+j}-{i}. a$k * b$(i+j-k))"
   12.70 -    by (rule setsum_diff1') simp_all
   12.71 +    by (rule setsum.remove) simp_all
   12.72    also have "(\<Sum>k\<in>{0..i+j}-{i}. a$k * b$(i+j-k)) = 0"
   12.73 -    proof (rule setsum_0' [rule_format])
   12.74 +    proof (rule setsum.neutral [rule_format])
   12.75        fix k assume "k \<in> {0..i+j} - {i}"
   12.76        then have "k < i \<or> i+j-k < j" by auto
   12.77        then show "a$k * b$(i+j-k) = 0" using i j by auto
   12.78 @@ -388,7 +388,7 @@
   12.79    have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
   12.80      by (simp add: fps_mult_nth)
   12.81    also have "\<dots> = f $ (n - 1)"
   12.82 -    using False by (simp add: X_def mult_delta_left setsum_delta)
   12.83 +    using False by (simp add: X_def mult_delta_left setsum.delta)
   12.84    finally show ?thesis using False by simp
   12.85  next
   12.86    case True
   12.87 @@ -577,7 +577,7 @@
   12.88  lemma fps_sum_rep_nth: "(setsum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
   12.89      (if n \<le> m then a$n else 0::'a::comm_ring_1)"
   12.90    apply (auto simp add: fps_setsum_nth cond_value_iff cong del: if_weak_cong)
   12.91 -  apply (simp add: setsum_delta')
   12.92 +  apply (simp add: setsum.delta')
   12.93    done
   12.94  
   12.95  lemma fps_notation: "(\<lambda>n. setsum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) ----> a"
   12.96 @@ -631,7 +631,7 @@
   12.97  
   12.98  subsection{* Inverses of formal power series *}
   12.99  
  12.100 -declare setsum_cong[fundef_cong]
  12.101 +declare setsum.cong[fundef_cong]
  12.102  
  12.103  instantiation fps :: ("{comm_monoid_add, inverse, times, uminus}") inverse
  12.104  begin
  12.105 @@ -756,10 +756,10 @@
  12.106    let ?g = "\<lambda>i. if i = n then 1 else if i=n - 1 then - 1 else 0"
  12.107    let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
  12.108    have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
  12.109 -    by (rule setsum_cong2) auto
  12.110 +    by (rule setsum.cong) auto
  12.111    have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
  12.112      apply (insert n)
  12.113 -    apply (rule setsum_cong2)
  12.114 +    apply (rule setsum.cong)
  12.115      apply auto
  12.116      done
  12.117    have eq: "{0 .. n} = {0.. n - 1} \<union> {n}"
  12.118 @@ -770,9 +770,9 @@
  12.119      by auto
  12.120    show "setsum ?f {0..n} = 0"
  12.121      unfolding th1
  12.122 -    apply (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
  12.123 +    apply (simp add: setsum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
  12.124      unfolding th2
  12.125 -    apply (simp add: setsum_delta)
  12.126 +    apply (simp add: setsum.delta)
  12.127      done
  12.128  qed
  12.129  
  12.130 @@ -811,15 +811,15 @@
  12.131      have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
  12.132        by (simp only: mult_commute)
  12.133      also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
  12.134 -      by (simp add: fps_mult_nth setsum_addf[symmetric])
  12.135 +      by (simp add: fps_mult_nth setsum.distrib[symmetric])
  12.136      also have "\<dots> = setsum ?h {0..n+1}"
  12.137        by (rule setsum.reindex_bij_witness_not_neutral
  12.138              [where S'="{}" and T'="{0}" and j="Suc" and i="\<lambda>i. i - 1"]) auto
  12.139      also have "\<dots> = (fps_deriv (f * g)) $ n"
  12.140 -      apply (simp only: fps_deriv_nth fps_mult_nth setsum_addf)
  12.141 +      apply (simp only: fps_deriv_nth fps_mult_nth setsum.distrib)
  12.142        unfolding s0 s1
  12.143 -      unfolding setsum_addf[symmetric] setsum_right_distrib
  12.144 -      apply (rule setsum_cong2)
  12.145 +      unfolding setsum.distrib[symmetric] setsum_right_distrib
  12.146 +      apply (rule setsum.cong)
  12.147        apply (auto simp add: of_nat_diff field_simps)
  12.148        done
  12.149      finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .
  12.150 @@ -1030,7 +1030,7 @@
  12.151          also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
  12.152            by (simp add: fps_mult_nth)
  12.153          also have "\<dots> = 0"
  12.154 -          apply (rule setsum_0')
  12.155 +          apply (rule setsum.neutral)
  12.156            apply auto
  12.157            apply (case_tac "x = m")
  12.158            using a0 apply simp
  12.159 @@ -1052,7 +1052,7 @@
  12.160    assumes a0: "a $0 = (0::'a::idom)"
  12.161      and kn: "n \<ge> k"
  12.162    shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
  12.163 -  apply (rule setsum_mono_zero_right)
  12.164 +  apply (rule setsum.mono_neutral_right)
  12.165    using kn
  12.166    apply auto
  12.167    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
  12.168 @@ -1072,7 +1072,7 @@
  12.169    also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
  12.170      by (simp add: fps_mult_nth)
  12.171    also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
  12.172 -    apply (rule setsum_mono_zero_right)
  12.173 +    apply (rule setsum.mono_neutral_right)
  12.174      apply simp
  12.175      apply clarsimp
  12.176      apply clarsimp
  12.177 @@ -1255,11 +1255,11 @@
  12.178    by (simp add: fps_compose_def)
  12.179  
  12.180  lemma fps_compose_X[simp]: "a oo X = (a :: 'a::comm_ring_1 fps)"
  12.181 -  by (simp add: fps_ext fps_compose_def mult_delta_right setsum_delta')
  12.182 +  by (simp add: fps_ext fps_compose_def mult_delta_right setsum.delta')
  12.183  
  12.184  lemma fps_const_compose[simp]:
  12.185    "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
  12.186 -  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
  12.187 +  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
  12.188  
  12.189  lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k"
  12.190    unfolding numeral_fps_const by simp
  12.191 @@ -1268,7 +1268,7 @@
  12.192    unfolding neg_numeral_fps_const by simp
  12.193  
  12.194  lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: 'a::comm_ring_1 fps)"
  12.195 -  by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum_delta not_le)
  12.196 +  by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum.delta not_le)
  12.197  
  12.198  
  12.199  subsection {* Rules from Herbert Wilf's Generatingfunctionology*}
  12.200 @@ -1374,10 +1374,10 @@
  12.201          by (simp add: fps_mult_nth)
  12.202        also have "\<dots> = a$n"
  12.203          unfolding th0
  12.204 -        unfolding setsum_Un_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
  12.205 -        unfolding setsum_Un_disjoint[OF f(2) f(3) d(2)]
  12.206 +        unfolding setsum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
  12.207 +        unfolding setsum.union_disjoint[OF f(2) f(3) d(2)]
  12.208          apply (simp)
  12.209 -        unfolding setsum_Un_disjoint[OF f(4,5) d(3), unfolded u(3)]
  12.210 +        unfolding setsum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
  12.211          apply simp
  12.212          done
  12.213        finally have "a$n = ((1 - X) * ?sa) $ n"
  12.214 @@ -1532,7 +1532,7 @@
  12.215        apply (auto simp add: atLeastLessThanSuc_atLeastAtMost listsum_setsum_nth)
  12.216        done
  12.217      also have "\<dots> = n + setsum (nth xs) ({0..k} - {i})"
  12.218 -      unfolding setsum_Un_disjoint[OF f d, unfolded eqs] using i by simp
  12.219 +      unfolding setsum.union_disjoint[OF f d, unfolded eqs] using i by simp
  12.220      finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
  12.221        by auto
  12.222      from H have xsl: "length xs = k+1"
  12.223 @@ -1565,8 +1565,8 @@
  12.224      have "listsum ?xs = setsum (nth ?xs) {0..<k+1}"
  12.225        unfolding listsum_setsum_nth xsl ..
  12.226      also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
  12.227 -      by (rule setsum_cong2) (simp del: replicate.simps)
  12.228 -    also have "\<dots> = n" using i by (simp add: setsum_delta)
  12.229 +      by (rule setsum.cong) (simp_all del: replicate.simps)
  12.230 +    also have "\<dots> = n" using i by (simp add: setsum.delta)
  12.231      finally have "?xs \<in> natpermute n (k+1)"
  12.232        using xsl unfolding natpermute_def mem_Collect_eq by blast
  12.233      then have "?xs \<in> ?A"
  12.234 @@ -1600,14 +1600,14 @@
  12.235      have f0: "finite {0 .. k}" "finite {m}" by auto
  12.236      have d0: "{0 .. k} \<inter> {m} = {}" using Suc by auto
  12.237      have "(setprod a {0 .. m}) $ n = (setprod a {0 .. k} * a m) $ n"
  12.238 -      unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0] by simp
  12.239 +      unfolding setprod.union_disjoint[OF f0 d0, unfolded u0] by simp
  12.240      also have "\<dots> = (\<Sum>i = 0..n. (\<Sum>v\<in>natpermute i (k + 1). \<Prod>j\<in>{0..k}. a j $ v ! j) * a m $ (n - i))"
  12.241        unfolding fps_mult_nth H[rule_format, OF km] ..
  12.242      also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
  12.243        apply (simp add: Suc)
  12.244        unfolding natpermute_split[of m "m + 1", simplified, of n,
  12.245          unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
  12.246 -      apply (subst setsum_UN_disjoint)
  12.247 +      apply (subst setsum.UNION_disjoint)
  12.248        apply simp
  12.249        apply simp
  12.250        unfolding image_Collect[symmetric]
  12.251 @@ -1616,13 +1616,14 @@
  12.252        apply (rule natpermute_finite)
  12.253        apply (clarsimp simp add: set_eq_iff)
  12.254        apply auto
  12.255 -      apply (rule setsum_cong2)
  12.256 +      apply (rule setsum.cong)
  12.257 +      apply (rule refl)
  12.258        unfolding setsum_left_distrib
  12.259        apply (rule sym)
  12.260 -      apply (rule_tac f="\<lambda>xs. xs @[n - x]" in  setsum_reindex_cong)
  12.261 +      apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in setsum.reindex_cong)
  12.262        apply (simp add: inj_on_def)
  12.263        apply auto
  12.264 -      unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
  12.265 +      unfolding setprod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
  12.266        apply (clarsimp simp add: natpermute_def nth_append)
  12.267        done
  12.268      finally show ?thesis .
  12.269 @@ -1692,16 +1693,16 @@
  12.270          have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using n1 by auto
  12.271          have d: "{0 .. n1} \<inter> {n} = {}" using n1 by auto
  12.272          have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
  12.273 -          apply (rule setsum_cong2)
  12.274 +          apply (rule setsum.cong)
  12.275            using H n1
  12.276            apply auto
  12.277            done
  12.278          have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
  12.279 -          unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq] seq
  12.280 +          unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq] seq
  12.281            using startsby_zero_power_nth_same[OF a0]
  12.282            by simp
  12.283          have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
  12.284 -          unfolding fps_compose_nth setsum_Un_disjoint[OF f d, unfolded eq]
  12.285 +          unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq]
  12.286            using startsby_zero_power_nth_same[OF a0]
  12.287            by simp
  12.288          from h[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
  12.289 @@ -1715,7 +1716,7 @@
  12.290  
  12.291  subsection {* Radicals *}
  12.292  
  12.293 -declare setprod_cong [fundef_cong]
  12.294 +declare setprod.cong [fundef_cong]
  12.295  
  12.296  function radical :: "(nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a::field fps \<Rightarrow> nat \<Rightarrow> 'a"
  12.297  where
  12.298 @@ -1752,8 +1753,8 @@
  12.299        also have "\<dots> = setsum (nth xs) {0..<Suc k}" using xs
  12.300          by (simp add: natpermute_def listsum_setsum_nth)
  12.301        also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
  12.302 -        unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
  12.303 -        unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
  12.304 +        unfolding eqs  setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
  12.305 +        unfolding setsum.union_disjoint[OF fths(2) fths(3) d(2)]
  12.306          by simp
  12.307        finally have False using c' by simp
  12.308      }
  12.309 @@ -1789,7 +1790,7 @@
  12.310    have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
  12.311      unfolding fps_power_nth Suc by simp
  12.312    also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
  12.313 -    apply (rule setprod_cong)
  12.314 +    apply (rule setprod.cong)
  12.315      apply simp
  12.316      using Suc
  12.317      apply (subgoal_tac "replicate k 0 ! x = 0")
  12.318 @@ -1865,7 +1866,7 @@
  12.319              by (metis natpermute_finite)+
  12.320            let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
  12.321            have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
  12.322 -          proof (rule setsum_cong2)
  12.323 +          proof (rule setsum.cong)
  12.324              fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
  12.325              let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
  12.326                fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
  12.327 @@ -1873,21 +1874,21 @@
  12.328                unfolding natpermute_contain_maximal by auto
  12.329              have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
  12.330                  (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
  12.331 -              apply (rule setprod_cong, simp)
  12.332 +              apply (rule setprod.cong, simp)
  12.333                using i r0
  12.334                apply (simp del: replicate.simps)
  12.335                done
  12.336              also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
  12.337                using i r0 by (simp add: setprod_gen_delta)
  12.338              finally show ?ths .
  12.339 -          qed
  12.340 +          qed rule
  12.341            then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
  12.342              by (simp add: natpermute_max_card[OF nz, simplified])
  12.343            also have "\<dots> = a$n - setsum ?f ?Pnknn"
  12.344              unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
  12.345            finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
  12.346            have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
  12.347 -            unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
  12.348 +            unfolding fps_power_nth_Suc setsum.union_disjoint[OF f d, unfolded eq] ..
  12.349            also have "\<dots> = a$n" unfolding fn by simp
  12.350            finally have "?r ^ Suc k $ n = a $n" .
  12.351          }
  12.352 @@ -1934,13 +1935,13 @@
  12.353            by (metis natpermute_finite)+
  12.354          let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
  12.355          have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
  12.356 -        proof(rule setsum_cong2)
  12.357 +        proof(rule setsum.cong2)
  12.358            fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
  12.359            let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
  12.360            from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
  12.361              unfolding natpermute_contain_maximal by auto
  12.362            have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
  12.363 -            apply (rule setprod_cong, simp)
  12.364 +            apply (rule setprod.cong, simp)
  12.365              using i r0 by (simp del: replicate.simps)
  12.366            also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
  12.367              unfolding setprod_gen_delta[OF fK] using i r0 by simp
  12.368 @@ -1952,7 +1953,7 @@
  12.369            unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
  12.370          finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
  12.371          have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
  12.372 -          unfolding fps_power_nth_Suc setsum_Un_disjoint[OF f d, unfolded eq] ..
  12.373 +          unfolding fps_power_nth_Suc setsum.union_disjoint[OF f d, unfolded eq] ..
  12.374          also have "\<dots> = a$n" unfolding fn by simp
  12.375          finally have "?r ^ Suc k $ n = a $n" .}
  12.376        ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
  12.377 @@ -2010,7 +2011,7 @@
  12.378          let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
  12.379          let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
  12.380          have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
  12.381 -        proof (rule setsum_cong2)
  12.382 +        proof (rule setsum.cong)
  12.383            fix v
  12.384            assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
  12.385            let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
  12.386 @@ -2018,18 +2019,18 @@
  12.387              unfolding Suc_eq_plus1 natpermute_contain_maximal
  12.388              by (auto simp del: replicate.simps)
  12.389            have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
  12.390 -            apply (rule setprod_cong, simp)
  12.391 +            apply (rule setprod.cong, simp)
  12.392              using i a0
  12.393              apply (simp del: replicate.simps)
  12.394              done
  12.395            also have "\<dots> = a $ n * (?r $ 0)^k"
  12.396              using i by (simp add: setprod_gen_delta)
  12.397            finally show ?ths .
  12.398 -        qed
  12.399 +        qed rule
  12.400          then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
  12.401            by (simp add: natpermute_max_card[OF nz, simplified])
  12.402          have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
  12.403 -        proof (rule setsum_cong2, rule setprod_cong, simp)
  12.404 +        proof (rule setsum.cong, rule refl, rule setprod.cong, simp)
  12.405            fix xs i
  12.406            assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
  12.407            {
  12.408 @@ -2048,8 +2049,8 @@
  12.409              also have "\<dots> = setsum (nth xs) {0..<Suc k}"
  12.410                using xs by (simp add: natpermute_def listsum_setsum_nth)
  12.411              also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
  12.412 -              unfolding eqs  setsum_Un_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
  12.413 -              unfolding setsum_Un_disjoint[OF fths(2) fths(3) d(2)]
  12.414 +              unfolding eqs  setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
  12.415 +              unfolding setsum.union_disjoint[OF fths(2) fths(3) d(2)]
  12.416                by simp
  12.417              finally have False using c' by simp
  12.418            }
  12.419 @@ -2062,7 +2063,7 @@
  12.420            by (simp add: fps_eq_iff)
  12.421          also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
  12.422            unfolding fps_power_nth_Suc
  12.423 -          using setsum_Un_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
  12.424 +          using setsum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
  12.425              unfolded eq, of ?g] by simp
  12.426          also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn"
  12.427            unfolding th0 th1 ..
  12.428 @@ -2283,12 +2284,12 @@
  12.429      also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
  12.430        unfolding fps_mult_nth ..
  12.431      also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
  12.432 -      apply (rule setsum_mono_zero_right)
  12.433 -      apply (auto simp add: mult_delta_left setsum_delta not_le)
  12.434 +      apply (rule setsum.mono_neutral_right)
  12.435 +      apply (auto simp add: mult_delta_left setsum.delta not_le)
  12.436        done
  12.437      also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
  12.438        unfolding fps_deriv_nth
  12.439 -      by (rule setsum_reindex_cong [where f = Suc]) (auto simp add: mult_assoc)
  12.440 +      by (rule setsum.reindex_cong [of Suc]) (auto simp add: mult_assoc)
  12.441      finally have th0: "(fps_deriv (a oo b))$n =
  12.442        setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
  12.443  
  12.444 @@ -2296,8 +2297,9 @@
  12.445        unfolding fps_mult_nth by (simp add: mult_ac)
  12.446      also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
  12.447        unfolding fps_deriv_nth fps_compose_nth setsum_right_distrib mult_assoc
  12.448 -      apply (rule setsum_cong2)
  12.449 -      apply (rule setsum_mono_zero_left)
  12.450 +      apply (rule setsum.cong)
  12.451 +      apply (rule refl)
  12.452 +      apply (rule setsum.mono_neutral_left)
  12.453        apply (simp_all add: subset_eq)
  12.454        apply clarify
  12.455        apply (subgoal_tac "b^i$x = 0")
  12.456 @@ -2307,8 +2309,8 @@
  12.457        done
  12.458      also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
  12.459        unfolding setsum_right_distrib
  12.460 -      apply (subst setsum_commute)
  12.461 -      apply (rule setsum_cong2)+
  12.462 +      apply (subst setsum.commute)
  12.463 +      apply (rule setsum.cong, rule refl)+
  12.464        apply simp
  12.465        done
  12.466      finally have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n"
  12.467 @@ -2328,7 +2330,7 @@
  12.468    have "((1+X)*a) $n = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0..n}"
  12.469      by (simp add: fps_mult_nth)
  12.470    also have "\<dots> = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
  12.471 -    unfolding Suc by (rule setsum_mono_zero_right) auto
  12.472 +    unfolding Suc by (rule setsum.mono_neutral_right) auto
  12.473    also have "\<dots> = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
  12.474      by (simp add: Suc)
  12.475    finally show ?thesis .
  12.476 @@ -2345,7 +2347,7 @@
  12.477      fix i
  12.478      have "a$i = ?r$i"
  12.479        unfolding fps_setsum_nth fps_mult_left_const_nth X_power_nth
  12.480 -      by (simp add: mult_delta_right setsum_delta' z)
  12.481 +      by (simp add: mult_delta_right setsum.delta' z)
  12.482    }
  12.483    then show ?thesis unfolding fps_eq_iff by blast
  12.484  qed
  12.485 @@ -2444,16 +2446,16 @@
  12.486    done
  12.487  
  12.488  lemma fps_compose_1[simp]: "1 oo a = 1"
  12.489 -  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
  12.490 +  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
  12.491  
  12.492  lemma fps_compose_0[simp]: "0 oo a = 0"
  12.493    by (simp add: fps_eq_iff fps_compose_nth)
  12.494  
  12.495  lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a$0)"
  12.496 -  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum_0')
  12.497 +  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum.neutral)
  12.498  
  12.499  lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
  12.500 -  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_addf)
  12.501 +  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum.distrib)
  12.502  
  12.503  lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
  12.504  proof (cases "finite S")
  12.505 @@ -2493,15 +2495,16 @@
  12.506      done
  12.507    have "?r =  setsum (\<lambda>i. setsum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
  12.508      apply (simp add: fps_mult_nth setsum_right_distrib)
  12.509 -    apply (subst setsum_commute)
  12.510 -    apply (rule setsum_cong2)
  12.511 +    apply (subst setsum.commute)
  12.512 +    apply (rule setsum.cong)
  12.513      apply (auto simp add: field_simps)
  12.514      done
  12.515    also have "\<dots> = ?l"
  12.516      apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
  12.517 -    apply (rule setsum_cong2)
  12.518 -    apply (simp add: setsum_cartesian_product mult_assoc)
  12.519 -    apply (rule setsum_mono_zero_right[OF f])
  12.520 +    apply (rule setsum.cong)
  12.521 +    apply (rule refl)
  12.522 +    apply (simp add: setsum.cartesian_product mult_assoc)
  12.523 +    apply (rule setsum.mono_neutral_right[OF f])
  12.524      apply (simp add: subset_eq)
  12.525      apply presburger
  12.526      apply clarsimp
  12.527 @@ -2524,8 +2527,8 @@
  12.528    shows "((a oo c) * (b oo d))$n =
  12.529      setsum (\<lambda>k. setsum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}"  (is "?l = ?r")
  12.530    unfolding product_composition_lemma[OF c0 d0]
  12.531 -  unfolding setsum_cartesian_product
  12.532 -  apply (rule setsum_mono_zero_left)
  12.533 +  unfolding setsum.cartesian_product
  12.534 +  apply (rule setsum.mono_neutral_left)
  12.535    apply simp
  12.536    apply (clarsimp simp add: subset_eq)
  12.537    apply clarsimp
  12.538 @@ -2533,7 +2536,7 @@
  12.539    apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
  12.540    apply simp
  12.541    unfolding fps_mult_nth
  12.542 -  apply (rule setsum_0')
  12.543 +  apply (rule setsum.neutral)
  12.544    apply (clarsimp simp add: not_le)
  12.545    apply (case_tac "x < aa")
  12.546    apply (rule startsby_zero_power_prefix[OF c0, rule_format])
  12.547 @@ -2558,9 +2561,9 @@
  12.548      done
  12.549    show "?l = ?r "
  12.550      unfolding th0
  12.551 -    apply (subst setsum_UN_disjoint)
  12.552 +    apply (subst setsum.UNION_disjoint)
  12.553      apply auto
  12.554 -    apply (subst setsum_UN_disjoint)
  12.555 +    apply (subst setsum.UNION_disjoint)
  12.556      apply auto
  12.557      done
  12.558  qed
  12.559 @@ -2613,7 +2616,7 @@
  12.560    using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
  12.561  
  12.562  lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
  12.563 -  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum_delta)
  12.564 +  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
  12.565  
  12.566  lemma fps_inverse_compose:
  12.567    assumes b0: "(b$0 :: 'a::field) = 0"
  12.568 @@ -2706,8 +2709,9 @@
  12.569        by (simp add: fps_compose_setsum_distrib)
  12.570      also have "\<dots> = ?r$n"
  12.571        apply (simp add: fps_compose_nth fps_setsum_nth setsum_left_distrib mult_assoc)
  12.572 -      apply (rule setsum_cong2)
  12.573 -      apply (rule setsum_mono_zero_right)
  12.574 +      apply (rule setsum.cong)
  12.575 +      apply (rule refl)
  12.576 +      apply (rule setsum.mono_neutral_right)
  12.577        apply (auto simp add: not_le)
  12.578        apply (erule startsby_zero_power_prefix[OF b0, rule_format])
  12.579        done
  12.580 @@ -2737,7 +2741,7 @@
  12.581      {
  12.582        assume kn: "k \<le> n"
  12.583        then have "?l$n = ?r$n"
  12.584 -        by (simp add: fps_compose_nth mult_delta_left setsum_delta)
  12.585 +        by (simp add: fps_compose_nth mult_delta_left setsum.delta)
  12.586      }
  12.587      moreover have "k >n \<or> k\<le> n"  by arith
  12.588      ultimately have "?l$n = ?r$n"  by blast
  12.589 @@ -2978,7 +2982,7 @@
  12.590  
  12.591  lemma Ec_E1_eq: "E (1::'a::field_char_0) oo (fps_const c * X) = E c"
  12.592    apply (auto simp add: fps_eq_iff E_def fps_compose_def power_mult_distrib)
  12.593 -  apply (simp add: cond_value_iff cond_application_beta setsum_delta' cong del: if_weak_cong)
  12.594 +  apply (simp add: cond_value_iff cond_application_beta setsum.delta' cong del: if_weak_cong)
  12.595    done
  12.596  
  12.597  text{* The generalized binomial theorem as a  consequence of @{thm E_add_mult} *}
  12.598 @@ -2994,8 +2998,8 @@
  12.599      by (simp add: field_simps fps_mult_nth of_nat_mult[symmetric] setsum_right_distrib)
  12.600    then show ?thesis
  12.601      apply simp
  12.602 -    apply (rule setsum_cong2)
  12.603 -    apply simp
  12.604 +    apply (rule setsum.cong)
  12.605 +    apply simp_all
  12.606      apply (frule binomial_fact[where ?'a = 'a, symmetric])
  12.607      apply (simp add: field_simps of_nat_mult)
  12.608      done
  12.609 @@ -3207,7 +3211,7 @@
  12.610    using binomial_Vandermonde[of n n n,symmetric]
  12.611    unfolding mult_2
  12.612    apply (simp add: power2_eq_square)
  12.613 -  apply (rule setsum_cong2)
  12.614 +  apply (rule setsum.cong)
  12.615    apply (auto intro:  binomial_symmetric)
  12.616    done
  12.617  
  12.618 @@ -3288,19 +3292,19 @@
  12.619            apply (simp add: field_simps del: fact_Suc)
  12.620            unfolding fact_altdef_nat id_def
  12.621            unfolding of_nat_setprod
  12.622 -          unfolding setprod_timesf[symmetric]
  12.623 +          unfolding setprod.distrib[symmetric]
  12.624            apply auto
  12.625            unfolding eq1
  12.626 -          apply (subst setprod_Un_disjoint[symmetric])
  12.627 +          apply (subst setprod.union_disjoint[symmetric])
  12.628            apply (auto)
  12.629 -          apply (rule setprod_cong)
  12.630 +          apply (rule setprod.cong)
  12.631            apply auto
  12.632            done
  12.633          have th20: "?m1 n * ?p b n = setprod (\<lambda>i. b - of_nat i) {0..m}"
  12.634            unfolding m1nk
  12.635            unfolding m h pochhammer_Suc_setprod
  12.636 -          unfolding setprod_timesf[symmetric]
  12.637 -          apply (rule setprod_cong)
  12.638 +          unfolding setprod.distrib[symmetric]
  12.639 +          apply (rule setprod.cong)
  12.640            apply auto
  12.641            done
  12.642          have th21:"pochhammer (b - of_nat n + 1) k = setprod (\<lambda>i. b - of_nat i) {n - k .. n - 1}"
  12.643 @@ -3314,10 +3318,10 @@
  12.644            pochhammer (b - of_nat n + 1) k * setprod (\<lambda>i. b - of_nat i) {0.. n - k - 1}"
  12.645            unfolding th20 th21
  12.646            unfolding h m
  12.647 -          apply (subst setprod_Un_disjoint[symmetric])
  12.648 +          apply (subst setprod.union_disjoint[symmetric])
  12.649            using kn' h m
  12.650            apply auto
  12.651 -          apply (rule setprod_cong)
  12.652 +          apply (rule setprod.cong)
  12.653            apply auto
  12.654            done
  12.655          then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k =
  12.656 @@ -3359,7 +3363,8 @@
  12.657      unfolding gbinomial_pochhammer
  12.658      using bn0
  12.659      apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
  12.660 -    apply (rule setsum_cong2)
  12.661 +    apply (rule setsum.cong)
  12.662 +    apply (rule refl)
  12.663      apply (drule th00(2))
  12.664      apply (simp add: field_simps power_add[symmetric])
  12.665      done
  12.666 @@ -3688,7 +3693,7 @@
  12.667    have th0: "(fps_const c * X) $ 0 = 0" by simp
  12.668    show ?thesis unfolding gp[OF th0, symmetric]
  12.669      by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
  12.670 -      fps_compose_nth power_mult_distrib cond_value_iff setsum_delta' cong del: if_weak_cong)
  12.671 +      fps_compose_nth power_mult_distrib cond_value_iff setsum.delta' cong del: if_weak_cong)
  12.672  qed
  12.673  
  12.674  lemma F_B[simp]: "F [-a] [] (- 1) = fps_binomial a"
  12.675 @@ -3750,7 +3755,7 @@
  12.676  
  12.677  lemma setsum_eq_if: "setsum f {(n::nat) .. m} = (if m < n then 0 else f n + setsum f {n+1 .. m})"
  12.678    apply simp
  12.679 -  apply (subst setsum_insert[symmetric])
  12.680 +  apply (subst setsum.insert[symmetric])
  12.681    apply (auto simp add: not_less setsum_head_Suc)
  12.682    done
  12.683  
    13.1 --- a/src/HOL/Library/Indicator_Function.thy	Fri Jun 27 22:08:55 2014 +0200
    13.2 +++ b/src/HOL/Library/Indicator_Function.thy	Sat Jun 28 09:16:42 2014 +0200
    13.3 @@ -57,7 +57,7 @@
    13.4    shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator B x) = (\<Sum>x \<in> A \<inter> B. f x)"
    13.5    and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator B x * f x) = (\<Sum>x \<in> A \<inter> B. f x)"
    13.6    unfolding indicator_def
    13.7 -  using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm)
    13.8 +  using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm)
    13.9  
   13.10  lemma setsum_indicator_eq_card:
   13.11    assumes "finite A"
   13.12 @@ -68,6 +68,6 @@
   13.13  lemma setsum_indicator_scaleR[simp]:
   13.14    "finite A \<Longrightarrow>
   13.15      (\<Sum>x \<in> A. indicator (B x) (g x) *\<^sub>R f x) = (\<Sum>x \<in> {x\<in>A. g x \<in> B x}. f x::'a::real_vector)"
   13.16 -  using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm simp: indicator_def)
   13.17 +  using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm simp: indicator_def)
   13.18  
   13.19  end
   13.20 \ No newline at end of file
    14.1 --- a/src/HOL/Library/Multiset.thy	Fri Jun 27 22:08:55 2014 +0200
    14.2 +++ b/src/HOL/Library/Multiset.thy	Sat Jun 28 09:16:42 2014 +0200
    14.3 @@ -612,7 +612,7 @@
    14.4  
    14.5  lemma size_multiset_union [simp]:
    14.6    "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N"
    14.7 -apply (simp add: size_multiset_def setsum_Un_nat setsum_addf setsum_wcount_Int wcount_union)
    14.8 +apply (simp add: size_multiset_def setsum_Un_nat setsum.distrib setsum_wcount_Int wcount_union)
    14.9  apply (subst Int_commute)
   14.10  apply (simp add: setsum_wcount_Int)
   14.11  done
   14.12 @@ -2287,10 +2287,10 @@
   14.13    hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
   14.14    hence A: "?A = \<Union> {?As b | b.  b \<in> ?B}" by auto
   14.15    have "setsum (count f) ?A = setsum (setsum (count f)) {?As b | b.  b \<in> ?B}"
   14.16 -    unfolding A by transfer (intro setsum_Union_disjoint, auto simp: multiset_def)
   14.17 +    unfolding A by transfer (intro setsum.Union_disjoint [simplified], auto simp: multiset_def setsum.Union_disjoint)
   14.18    also have "... = setsum (setsum (count f)) (?As ` ?B)" unfolding 0 ..
   14.19    also have "... = setsum (setsum (count f) o ?As) ?B"
   14.20 -    by(intro setsum_reindex) (auto simp add: setsum_gt_0_iff inj_on_def)
   14.21 +    by (intro setsum.reindex) (auto simp add: setsum_gt_0_iff inj_on_def)
   14.22    also have "... = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" unfolding comp_def ..
   14.23    finally have "setsum (count f) ?A = setsum (\<lambda> b. setsum (count f) (?As b)) ?B" .
   14.24    thus "count (mmap (h2 \<circ> h1) f) c = count ((mmap h2 \<circ> mmap h1) f) c"
   14.25 @@ -2300,7 +2300,7 @@
   14.26  lemma mmap_cong:
   14.27  assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a"
   14.28  shows "mmap f M = mmap g M"
   14.29 -using assms by transfer (auto intro!: setsum_cong)
   14.30 +using assms by transfer (auto intro!: setsum.cong)
   14.31  
   14.32  context
   14.33  begin
   14.34 @@ -2458,7 +2458,7 @@
   14.35    (*  *)
   14.36    let ?ct1 = "N1 o e1"  let ?ct2 = "N2 o e2"
   14.37    have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
   14.38 -  unfolding setsum_reindex[OF e1_inj, symmetric] setsum_reindex[OF e2_inj, symmetric]
   14.39 +  unfolding setsum.reindex[OF e1_inj, symmetric] setsum.reindex[OF e2_inj, symmetric]
   14.40    e1_surj e2_surj using ss .
   14.41    obtain ct where
   14.42    ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
   14.43 @@ -2490,8 +2490,8 @@
   14.44      hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def
   14.45      by (metis image_eqI lessThan_iff less_Suc_eq_le)
   14.46      have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))"
   14.47 -    unfolding e2_surj[symmetric] setsum_reindex[OF e2_inj]
   14.48 -    unfolding M_def comp_def apply(intro setsum_cong) apply force
   14.49 +    unfolding e2_surj[symmetric] setsum.reindex[OF e2_inj]
   14.50 +    unfolding M_def comp_def apply(intro setsum.cong) apply force
   14.51      by (metis e2_surj b1 h1 h2 imageI)
   14.52      also have "... = N1 b1" using b1 ct1[OF f1b1] by simp
   14.53      finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" .
   14.54 @@ -2500,8 +2500,8 @@
   14.55      hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def
   14.56      by (metis image_eqI lessThan_iff less_Suc_eq_le)
   14.57      have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))"
   14.58 -    unfolding e1_surj[symmetric] setsum_reindex[OF e1_inj]
   14.59 -    unfolding M_def comp_def apply(intro setsum_cong) apply force
   14.60 +    unfolding e1_surj[symmetric] setsum.reindex[OF e1_inj]
   14.61 +    unfolding M_def comp_def apply(intro setsum.cong) apply force
   14.62      by (metis e1_surj b2 h1 h2 imageI)
   14.63      also have "... = N2 b2" using b2 ct2[OF f2b2] by simp
   14.64      finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" .
   14.65 @@ -2610,7 +2610,7 @@
   14.66    have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
   14.67      unfolding set1_def by auto
   14.68    have setsum_set1: "\<And> c. setsum (count N1) (set1 c) = count P c"
   14.69 -    unfolding P1 set1_def by transfer (auto intro: setsum_cong)
   14.70 +    unfolding P1 set1_def by transfer (auto intro: setsum.cong)
   14.71  
   14.72    def set2 \<equiv> "\<lambda> c. {b2 \<in> set_of N2. f2 b2 = c}"
   14.73    have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto
   14.74 @@ -2625,7 +2625,7 @@
   14.75    have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
   14.76      unfolding set2_def by auto
   14.77    have setsum_set2: "\<And> c. setsum (count N2) (set2 c) = count P c"
   14.78 -    unfolding P2 set2_def by transfer (auto intro: setsum_cong)
   14.79 +    unfolding P2 set2_def by transfer (auto intro: setsum.cong)
   14.80  
   14.81    have ss: "\<And> c. c \<in> set_of P \<Longrightarrow> setsum (count N1) (set1 c) = setsum (count N2) (set2 c)"
   14.82      unfolding setsum_set1 setsum_set2 ..
   14.83 @@ -2746,9 +2746,9 @@
   14.84          have c: "c \<in> set_of P" and b1: "b1 \<in> set1 c"
   14.85            unfolding set1_def c_def P1 using True by (auto simp: comp_eq_dest[OF set_of_mmap])
   14.86          with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p1 a = b1 \<and> a \<in> SET}"
   14.87 -          by transfer (force intro: setsum_mono_zero_cong_left split: split_if_asm)
   14.88 +          by transfer (force intro: setsum.mono_neutral_cong_left split: split_if_asm)
   14.89          also have "... = setsum (count M) ((\<lambda> b2. u c b1 b2) ` (set2 c))"
   14.90 -          apply(rule setsum_cong) using c b1 proof safe
   14.91 +          apply(rule setsum.cong) using c b1 proof safe
   14.92            fix a assume p1a: "p1 a \<in> set1 c" and "c \<in> set_of P" and "a \<in> SET"
   14.93            hence ac: "a \<in> sset c" using p1_rev by auto
   14.94            hence "a = u c (p1 a) (p2 a)" using c by auto
   14.95 @@ -2756,10 +2756,10 @@
   14.96            ultimately show "a \<in> u c (p1 a) ` set2 c" by auto
   14.97          qed auto
   14.98          also have "... = setsum (\<lambda> b2. count M (u c b1 b2)) (set2 c)"
   14.99 -          unfolding comp_def[symmetric] apply(rule setsum_reindex)
  14.100 +          unfolding comp_def[symmetric] apply(rule setsum.reindex)
  14.101            using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast
  14.102          also have "... = count N1 b1" unfolding ss1[OF c b1, symmetric]
  14.103 -          apply(rule setsum_cong[OF refl]) apply (transfer fixing: Ms u c b1 set2)
  14.104 +          apply(rule setsum.cong[OF refl]) apply (transfer fixing: Ms u c b1 set2)
  14.105            using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce
  14.106          finally show ?thesis .
  14.107        qed
  14.108 @@ -2781,9 +2781,9 @@
  14.109          have c: "c \<in> set_of P" and b2: "b2 \<in> set2 c"
  14.110            unfolding set2_def c_def P2 using True by (auto simp: comp_eq_dest[OF set_of_mmap])
  14.111          with sM(1) have "setsum (count M) ?K = setsum (count M) {a. p2 a = b2 \<and> a \<in> SET}"
  14.112 -          by transfer (force intro: setsum_mono_zero_cong_left split: split_if_asm)
  14.113 +          by transfer (force intro: setsum.mono_neutral_cong_left split: split_if_asm)
  14.114          also have "... = setsum (count M) ((\<lambda> b1. u c b1 b2) ` (set1 c))"
  14.115 -          apply(rule setsum_cong) using c b2 proof safe
  14.116 +          apply(rule setsum.cong) using c b2 proof safe
  14.117            fix a assume p2a: "p2 a \<in> set2 c" and "c \<in> set_of P" and "a \<in> SET"
  14.118            hence ac: "a \<in> sset c" using p2_rev by auto
  14.119            hence "a = u c (p1 a) (p2 a)" using c by auto
  14.120 @@ -2791,11 +2791,11 @@
  14.121            ultimately show "a \<in> (\<lambda>x. u c x (p2 a)) ` set1 c" by auto
  14.122          qed auto
  14.123          also have "... = setsum (count M o (\<lambda> b1. u c b1 b2)) (set1 c)"
  14.124 -          apply(rule setsum_reindex)
  14.125 +          apply(rule setsum.reindex)
  14.126            using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast
  14.127          also have "... = setsum (\<lambda> b1. count M (u c b1 b2)) (set1 c)" by simp
  14.128          also have "... = count N2 b2" unfolding ss2[OF c b2, symmetric] comp_def
  14.129 -          apply(rule setsum_cong[OF refl]) apply (transfer fixing: Ms u c b2 set1)
  14.130 +          apply(rule setsum.cong[OF refl]) apply (transfer fixing: Ms u c b2 set1)
  14.131            using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] set1_def by fastforce
  14.132          finally show ?thesis .
  14.133        qed
  14.134 @@ -2855,7 +2855,7 @@
  14.135    assume "M1 \<in> multiset" "M2 \<in> multiset"
  14.136    hence "setsum M1 {a. f a = x \<and> 0 < M1 a} = setsum M1 {a. f a = x \<and> 0 < M1 a + M2 a}"
  14.137          "setsum M2 {a. f a = x \<and> 0 < M2 a} = setsum M2 {a. f a = x \<and> 0 < M1 a + M2 a}"
  14.138 -    by (auto simp: multiset_def intro!: setsum_mono_zero_cong_left)
  14.139 +    by (auto simp: multiset_def intro!: setsum.mono_neutral_cong_left)
  14.140    then show "(\<Sum>a | f a = x \<and> 0 < M1 a + M2 a. M1 a + M2 a) =
  14.141         setsum M1 {a. f a = x \<and> 0 < M1 a} +
  14.142         setsum M2 {a. f a = x \<and> 0 < M2 a}"
  14.143 @@ -2897,7 +2897,7 @@
  14.144      by (metis (lifting, full_types) mem_Collect_eq neq0_conv setsum.neutral)
  14.145    have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)"
  14.146    apply safe
  14.147 -    apply (metis less_not_refl setsum_gt_0_iff setsum_infinite)
  14.148 +    apply (metis less_not_refl setsum_gt_0_iff setsum.infinite)
  14.149      by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff)
  14.150    hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
  14.151  
    15.1 --- a/src/HOL/Library/Permutations.thy	Fri Jun 27 22:08:55 2014 +0200
    15.2 +++ b/src/HOL/Library/Permutations.thy	Sat Jun 28 09:16:42 2014 +0200
    15.3 @@ -955,7 +955,7 @@
    15.4      using image_inverse_permutations by blast
    15.5    have th2: "?rhs = setsum (f \<circ> inv) ?S"
    15.6      by (simp add: o_def)
    15.7 -  from setsum_reindex[OF th0, of f] show ?thesis unfolding th1 th2 .
    15.8 +  from setsum.reindex[OF th0, of f] show ?thesis unfolding th1 th2 .
    15.9  qed
   15.10  
   15.11  lemma setum_permutations_compose_left:
   15.12 @@ -979,7 +979,7 @@
   15.13    qed
   15.14    have th3: "(op \<circ> q) ` ?S = ?S"
   15.15      using image_compose_permutations_left[OF q] by auto
   15.16 -  from setsum_reindex[OF th1, of f] show ?thesis unfolding th0 th1 th3 .
   15.17 +  from setsum.reindex[OF th1, of f] show ?thesis unfolding th0 th1 th3 .
   15.18  qed
   15.19  
   15.20  lemma sum_permutations_compose_right:
   15.21 @@ -1003,7 +1003,7 @@
   15.22    qed
   15.23    have th3: "(\<lambda>p. p \<circ> q) ` ?S = ?S"
   15.24      using image_compose_permutations_right[OF q] by auto
   15.25 -  from setsum_reindex[OF th1, of f]
   15.26 +  from setsum.reindex[OF th1, of f]
   15.27    show ?thesis unfolding th0 th1 th3 .
   15.28  qed
   15.29  
   15.30 @@ -1024,10 +1024,10 @@
   15.31      by blast
   15.32    show ?thesis
   15.33      unfolding permutes_insert
   15.34 -    unfolding setsum_cartesian_product
   15.35 +    unfolding setsum.cartesian_product
   15.36      unfolding th1[symmetric]
   15.37      unfolding th0
   15.38 -  proof (rule setsum_reindex)
   15.39 +  proof (rule setsum.reindex)
   15.40      let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"
   15.41      let ?P = "{p. p permutes S}"
   15.42      {
    16.1 --- a/src/HOL/List.thy	Fri Jun 27 22:08:55 2014 +0200
    16.2 +++ b/src/HOL/List.thy	Sat Jun 28 09:16:42 2014 +0200
    16.3 @@ -4525,7 +4525,7 @@
    16.4    moreover
    16.5    have "Suc (card A - Suc k) = card A - k" using Suc.prems by simp
    16.6    then have "(card A - k) * \<Prod>{Suc (card A - k)..card A} = \<Prod>{Suc (card A - Suc k)..card A}"
    16.7 -    by (subst setprod_insert[symmetric]) (simp add: atLeastAtMost_insertL)+
    16.8 +    by (subst setprod.insert[symmetric]) (simp add: atLeastAtMost_insertL)+
    16.9    ultimately show ?case
   16.10      by (simp add: card_image inj_Cons card_UN_disjoint Suc.hyps algebra_simps)
   16.11  qed
    17.1 --- a/src/HOL/MacLaurin.thy	Fri Jun 27 22:08:55 2014 +0200
    17.2 +++ b/src/HOL/MacLaurin.thy	Sat Jun 28 09:16:42 2014 +0200
    17.3 @@ -87,7 +87,7 @@
    17.4        + (B * (t^n / real(fact n)))))"
    17.5  
    17.6    have g2: "g 0 = 0 & g h = 0"
    17.7 -    by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum_reindex)
    17.8 +    by (simp add: m f_h g_def lessThan_Suc_eq_insert_0 image_iff diff_0 setsum.reindex)
    17.9  
   17.10    def difg \<equiv> "(%m t. diff m t -
   17.11      (setsum (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) {..<n-m}
   17.12 @@ -101,7 +101,7 @@
   17.13      using diff_Suc m unfolding difg_def by (rule Maclaurin_lemma2)
   17.14  
   17.15    have difg_eq_0: "\<forall>m<n. difg m 0 = 0"
   17.16 -    by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum_reindex)
   17.17 +    by (auto simp: difg_def m Suc_diff_le lessThan_Suc_eq_insert_0 image_iff setsum.reindex)
   17.18  
   17.19    have isCont_difg: "\<And>m x. \<lbrakk>m < n; 0 \<le> x; x \<le> h\<rbrakk> \<Longrightarrow> isCont (difg m) x"
   17.20      by (rule DERIV_isCont [OF difg_Suc [rule_format]]) simp
   17.21 @@ -225,7 +225,7 @@
   17.22      by (auto simp add: power_mult_distrib[symmetric])
   17.23    moreover
   17.24    have "(SUM m<n. -1 ^ m * diff m 0 * (- h) ^ m / real (fact m)) = (SUM m<n. diff m 0 * h ^ m / real (fact m))"
   17.25 -    by (auto intro: setsum_cong simp add: power_mult_distrib[symmetric])
   17.26 +    by (auto intro: setsum.cong simp add: power_mult_distrib[symmetric])
   17.27    ultimately have " h < - t \<and>
   17.28      - t < 0 \<and>
   17.29      f h =
   17.30 @@ -419,13 +419,13 @@
   17.31  apply (simp (no_asm))
   17.32  apply (simp (no_asm) add: sin_expansion_lemma)
   17.33  apply (force intro!: derivative_eq_intros)
   17.34 -apply (subst (asm) setsum_0', clarify, case_tac "x", simp, simp)
   17.35 +apply (subst (asm) setsum.neutral, clarify, case_tac "x", simp, simp)
   17.36  apply (cases n, simp, simp)
   17.37  apply (rule ccontr, simp)
   17.38  apply (drule_tac x = x in spec, simp)
   17.39  apply (erule ssubst)
   17.40  apply (rule_tac x = t in exI, simp)
   17.41 -apply (rule setsum_cong[OF refl])
   17.42 +apply (rule setsum.cong[OF refl])
   17.43  apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
   17.44  done
   17.45  
   17.46 @@ -450,7 +450,7 @@
   17.47  apply (force intro!: derivative_eq_intros)
   17.48  apply (erule ssubst)
   17.49  apply (rule_tac x = t in exI, simp)
   17.50 -apply (rule setsum_cong[OF refl])
   17.51 +apply (rule setsum.cong[OF refl])
   17.52  apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
   17.53  done
   17.54  
   17.55 @@ -467,7 +467,7 @@
   17.56  apply (force intro!: derivative_eq_intros)
   17.57  apply (erule ssubst)
   17.58  apply (rule_tac x = t in exI, simp)
   17.59 -apply (rule setsum_cong[OF refl])
   17.60 +apply (rule setsum.cong[OF refl])
   17.61  apply (auto simp add: sin_coeff_def sin_zero_iff odd_Suc_mult_two_ex)
   17.62  done
   17.63  
   17.64 @@ -497,7 +497,7 @@
   17.65  apply (drule_tac x = x in spec, simp)
   17.66  apply (erule ssubst)
   17.67  apply (rule_tac x = t in exI, simp)
   17.68 -apply (rule setsum_cong[OF refl])
   17.69 +apply (rule setsum.cong[OF refl])
   17.70  apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
   17.71  done
   17.72  
   17.73 @@ -513,7 +513,7 @@
   17.74  apply (simp (no_asm) add: cos_expansion_lemma)
   17.75  apply (erule ssubst)
   17.76  apply (rule_tac x = t in exI, simp)
   17.77 -apply (rule setsum_cong[OF refl])
   17.78 +apply (rule setsum.cong[OF refl])
   17.79  apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
   17.80  done
   17.81  
   17.82 @@ -529,7 +529,7 @@
   17.83  apply (simp (no_asm) add: cos_expansion_lemma)
   17.84  apply (erule ssubst)
   17.85  apply (rule_tac x = t in exI, simp)
   17.86 -apply (rule setsum_cong[OF refl])
   17.87 +apply (rule setsum.cong[OF refl])
   17.88  apply (auto simp add: cos_coeff_def cos_zero_iff even_mult_two_ex)
   17.89  done
   17.90  
   17.91 @@ -572,7 +572,7 @@
   17.92      unfolding sin_coeff_def
   17.93      apply (subst t2)
   17.94      apply (rule sin_bound_lemma)
   17.95 -    apply (rule setsum_cong[OF refl])
   17.96 +    apply (rule setsum.cong[OF refl])
   17.97      apply (subst diff_m_0, simp)
   17.98      apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
   17.99                  simp add: est mult_ac divide_inverse power_abs [symmetric] abs_mult)
    18.1 --- a/src/HOL/Metis_Examples/Big_O.thy	Fri Jun 27 22:08:55 2014 +0200
    18.2 +++ b/src/HOL/Metis_Examples/Big_O.thy	Sat Jun 28 09:16:42 2014 +0200
    18.3 @@ -552,7 +552,8 @@
    18.4   apply (erule ssubst)
    18.5   apply (erule bigo_setsum3)
    18.6  apply (rule ext)
    18.7 -apply (rule setsum_cong2)
    18.8 +apply (rule setsum.cong)
    18.9 +apply (rule refl)
   18.10  by (metis abs_of_nonneg zero_le_mult_iff)
   18.11  
   18.12  lemma bigo_setsum6: "f =o g +o O(h) \<Longrightarrow> \<forall>x y. 0 <= l x y \<Longrightarrow>
    19.1 --- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Fri Jun 27 22:08:55 2014 +0200
    19.2 +++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Sat Jun 28 09:16:42 2014 +0200
    19.3 @@ -64,13 +64,13 @@
    19.4    apply auto
    19.5    done
    19.6  
    19.7 -lemma setsum_Un_disjoint':
    19.8 +lemma setsum_union_disjoint':
    19.9    assumes "finite A"
   19.10      and "finite B"
   19.11      and "A \<inter> B = {}"
   19.12      and "A \<union> B = C"
   19.13    shows "setsum g C = setsum g A + setsum g B"
   19.14 -  using setsum_Un_disjoint[OF assms(1-3)] and assms(4) by auto
   19.15 +  using setsum.union_disjoint[OF assms(1-3)] and assms(4) by auto
   19.16  
   19.17  lemma pointwise_minimal_pointwise_maximal:
   19.18    fixes s :: "(nat \<Rightarrow> nat) set"
   19.19 @@ -151,19 +151,19 @@
   19.20    shows "odd (card {s\<in>S. compo s})"
   19.21  proof -
   19.22    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)"
   19.23 -    by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)
   19.24 +    by (subst setsum.union_disjoint[symmetric]) (auto intro!: setsum.cong)
   19.25    also have "\<dots> = (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> bnd f}. face f s}) +
   19.26                    (\<Sum>s\<in>S. card {f \<in> {f\<in>F. compo' f \<and> \<not> bnd f}. face f s})"
   19.27 -    unfolding setsum_addf[symmetric]
   19.28 +    unfolding setsum.distrib[symmetric]
   19.29      by (subst card_Un_disjoint[symmetric])
   19.30 -       (auto simp: nF_def intro!: setsum_cong arg_cong[where f=card])
   19.31 +       (auto simp: nF_def intro!: setsum.cong arg_cong[where f=card])
   19.32    also have "\<dots> = 1 * card {f\<in>F. compo' f \<and> bnd f} + 2 * card {f\<in>F. compo' f \<and> \<not> bnd f}"
   19.33      using assms(4,5) by (fastforce intro!: arg_cong2[where f="op +"] setsum_multicount)
   19.34    finally have "odd ((\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) + card {s\<in>S. compo s})"
   19.35      using assms(6,8) by simp
   19.36    moreover have "(\<Sum>s | s \<in> S \<and> \<not> compo s. nF s) =
   19.37      (\<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)"
   19.38 -    using assms(7) by (subst setsum_Un_disjoint[symmetric]) (fastforce intro!: setsum_cong)+
   19.39 +    using assms(7) by (subst setsum.union_disjoint[symmetric]) (fastforce intro!: setsum.cong)+
   19.40    ultimately show ?thesis
   19.41      by auto
   19.42  qed
    20.1 --- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    20.2 +++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    20.3 @@ -8,23 +8,18 @@
    20.4    "(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)"
    20.5    by (cases "k=a") auto
    20.6  
    20.7 -lemma setsum_Plus:
    20.8 -  "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow>
    20.9 -    (\<Sum>x\<in>A <+> B. g x) = (\<Sum>x\<in>A. g (Inl x)) + (\<Sum>x\<in>B. g (Inr x))"
   20.10 -  unfolding Plus_def
   20.11 -  by (subst setsum_Un_disjoint, auto simp add: setsum_reindex)
   20.12 -
   20.13  lemma setsum_UNIV_sum:
   20.14    fixes g :: "'a::finite + 'b::finite \<Rightarrow> _"
   20.15    shows "(\<Sum>x\<in>UNIV. g x) = (\<Sum>x\<in>UNIV. g (Inl x)) + (\<Sum>x\<in>UNIV. g (Inr x))"
   20.16    apply (subst UNIV_Plus_UNIV [symmetric])
   20.17 -  apply (rule setsum_Plus [OF finite finite])
   20.18 +  apply (subst setsum.Plus)
   20.19 +  apply simp_all
   20.20    done
   20.21  
   20.22  lemma setsum_mult_product:
   20.23    "setsum h {..<A * B :: nat} = (\<Sum>i\<in>{..<A}. \<Sum>j\<in>{..<B}. h (j + i * B))"
   20.24    unfolding setsum_nat_group[of h B A, unfolded atLeast0LessThan, symmetric]
   20.25 -proof (rule setsum_cong, simp, rule setsum_reindex_cong)
   20.26 +proof (rule setsum.cong, simp, rule setsum.reindex_cong)
   20.27    fix i
   20.28    show "inj_on (\<lambda>j. j + i * B) {..<B}" by (auto intro!: inj_onI)
   20.29    show "{i * B..<i * B + B} = (\<lambda>j. j + i * B) ` {..<B}"
   20.30 @@ -110,11 +105,17 @@
   20.31  
   20.32  subsection {* A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space. *}
   20.33  
   20.34 +lemma setsum_cong_aux:
   20.35 +  "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
   20.36 +  by (auto intro: setsum.cong)
   20.37 +
   20.38 +hide_fact (open) setsum_cong_aux
   20.39 +
   20.40  method_setup vector = {*
   20.41  let
   20.42    val ss1 =
   20.43      simpset_of (put_simpset HOL_basic_ss @{context}
   20.44 -      addsimps [@{thm setsum_addf} RS sym,
   20.45 +      addsimps [@{thm setsum.distrib} RS sym,
   20.46        @{thm setsum_subtractf} RS sym, @{thm setsum_right_distrib},
   20.47        @{thm setsum_left_distrib}, @{thm setsum_negf} RS sym])
   20.48    val ss2 =
   20.49 @@ -126,8 +127,8 @@
   20.50                @{thm vec_lambda_beta}, @{thm vector_scalar_mult_def}])
   20.51    fun vector_arith_tac ctxt ths =
   20.52      simp_tac (put_simpset ss1 ctxt)
   20.53 -    THEN' (fn i => rtac @{thm setsum_cong2} i
   20.54 -         ORELSE rtac @{thm setsum_0'} i
   20.55 +    THEN' (fn i => rtac @{thm Cartesian_Euclidean_Space.setsum_cong_aux} i
   20.56 +         ORELSE rtac @{thm setsum.neutral} i
   20.57           ORELSE simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm vec_eq_iff}]) i)
   20.58      (* THEN' TRY o clarify_tac HOL_cs  THEN' (TRY o rtac @{thm iffI}) *)
   20.59      THEN' asm_full_simp_tac (put_simpset ss2 ctxt addsimps ths)
   20.60 @@ -357,14 +358,14 @@
   20.61  
   20.62  lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def)
   20.63  lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)"
   20.64 -  by (vector matrix_matrix_mult_def setsum_addf[symmetric] field_simps)
   20.65 +  by (vector matrix_matrix_mult_def setsum.distrib[symmetric] field_simps)
   20.66  
   20.67  lemma matrix_mul_lid:
   20.68    fixes A :: "'a::semiring_1 ^ 'm ^ 'n"
   20.69    shows "mat 1 ** A = A"
   20.70    apply (simp add: matrix_matrix_mult_def mat_def)
   20.71    apply vector
   20.72 -  apply (auto simp only: if_distrib cond_application_beta setsum_delta'[OF finite]
   20.73 +  apply (auto simp only: if_distrib cond_application_beta setsum.delta'[OF finite]
   20.74      mult_1_left mult_zero_left if_True UNIV_I)
   20.75    done
   20.76  
   20.77 @@ -374,26 +375,26 @@
   20.78    shows "A ** mat 1 = A"
   20.79    apply (simp add: matrix_matrix_mult_def mat_def)
   20.80    apply vector
   20.81 -  apply (auto simp only: if_distrib cond_application_beta setsum_delta[OF finite]
   20.82 +  apply (auto simp only: if_distrib cond_application_beta setsum.delta[OF finite]
   20.83      mult_1_right mult_zero_right if_True UNIV_I cong: if_cong)
   20.84    done
   20.85  
   20.86  lemma matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C"
   20.87    apply (vector matrix_matrix_mult_def setsum_right_distrib setsum_left_distrib mult_assoc)
   20.88 -  apply (subst setsum_commute)
   20.89 +  apply (subst setsum.commute)
   20.90    apply simp
   20.91    done
   20.92  
   20.93  lemma matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x"
   20.94    apply (vector matrix_matrix_mult_def matrix_vector_mult_def
   20.95      setsum_right_distrib setsum_left_distrib mult_assoc)
   20.96 -  apply (subst setsum_commute)
   20.97 +  apply (subst setsum.commute)
   20.98    apply simp
   20.99    done
  20.100  
  20.101  lemma matrix_vector_mul_lid: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)"
  20.102    apply (vector matrix_vector_mult_def mat_def)
  20.103 -  apply (simp add: if_distrib cond_application_beta setsum_delta' cong del: if_weak_cong)
  20.104 +  apply (simp add: if_distrib cond_application_beta setsum.delta' cong del: if_weak_cong)
  20.105    done
  20.106  
  20.107  lemma matrix_transpose_mul:
  20.108 @@ -410,7 +411,7 @@
  20.109    apply (erule_tac x="axis ia 1" in allE)
  20.110    apply (erule_tac x="i" in allE)
  20.111    apply (auto simp add: if_distrib cond_application_beta axis_def
  20.112 -    setsum_delta[OF finite] cong del: if_weak_cong)
  20.113 +    setsum.delta[OF finite] cong del: if_weak_cong)
  20.114    done
  20.115  
  20.116  lemma matrix_vector_mul_component: "((A::real^_^_) *v x)$k = (A$k) \<bullet> x"
  20.117 @@ -418,7 +419,7 @@
  20.118  
  20.119  lemma dot_lmul_matrix: "((x::real ^_) v* A) \<bullet> y = x \<bullet> (A *v y)"
  20.120    apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def setsum_left_distrib setsum_right_distrib mult_ac)
  20.121 -  apply (subst setsum_commute)
  20.122 +  apply (subst setsum.commute)
  20.123    apply simp
  20.124    done
  20.125  
  20.126 @@ -455,10 +456,10 @@
  20.127  
  20.128  lemma vector_componentwise:
  20.129    "(x::'a::ring_1^'n) = (\<chi> j. \<Sum>i\<in>UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)"
  20.130 -  by (simp add: axis_def if_distrib setsum_cases vec_eq_iff)
  20.131 +  by (simp add: axis_def if_distrib setsum.If_cases vec_eq_iff)
  20.132  
  20.133  lemma basis_expansion: "setsum (\<lambda>i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)"
  20.134 -  by (auto simp add: axis_def vec_eq_iff if_distrib setsum_cases cong del: if_weak_cong)
  20.135 +  by (auto simp add: axis_def vec_eq_iff if_distrib setsum.If_cases cong del: if_weak_cong)
  20.136  
  20.137  lemma linear_componentwise:
  20.138    fixes f:: "real ^'m \<Rightarrow> real ^ _"
  20.139 @@ -492,7 +493,7 @@
  20.140  
  20.141  lemma matrix_vector_mul_linear: "linear(\<lambda>x. A *v (x::real ^ _))"
  20.142    by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff
  20.143 -      field_simps setsum_right_distrib setsum_addf)
  20.144 +      field_simps setsum_right_distrib setsum.distrib)
  20.145  
  20.146  lemma matrix_works:
  20.147    assumes lf: "linear f"
  20.148 @@ -523,7 +524,7 @@
  20.149    apply (rule adjoint_unique)
  20.150    apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
  20.151      setsum_left_distrib setsum_right_distrib)
  20.152 -  apply (subst setsum_commute)
  20.153 +  apply (subst setsum.commute)
  20.154    apply (auto simp add: mult_ac)
  20.155    done
  20.156  
  20.157 @@ -707,13 +708,13 @@
  20.158              using i(1) by (simp add: field_simps)
  20.159            have "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
  20.160                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"
  20.161 -            apply (rule setsum_cong[OF refl])
  20.162 +            apply (rule setsum.cong[OF refl])
  20.163              using th apply blast
  20.164              done
  20.165            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"
  20.166 -            by (simp add: setsum_addf)
  20.167 +            by (simp add: setsum.distrib)
  20.168            also have "\<dots> = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
  20.169 -            unfolding setsum_delta[OF fU]
  20.170 +            unfolding setsum.delta[OF fU]
  20.171              using i(1) by simp
  20.172            finally show "setsum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
  20.173              else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + setsum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
    21.1 --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    21.2 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    21.3 @@ -121,9 +121,9 @@
    21.4      by (auto intro: finite_subset[OF assms])
    21.5    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)"
    21.6      using d
    21.7 -    by (auto intro!: setsum_cong simp: inner_Basis inner_setsum_left)
    21.8 +    by (auto intro!: setsum.cong simp: inner_Basis inner_setsum_left)
    21.9    show ?thesis
   21.10 -    unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum_delta[OF **] ***)
   21.11 +    unfolding euclidean_eq_iff[where 'a='a] by (auto simp: setsum.delta[OF **] ***)
   21.12  qed
   21.13  
   21.14  lemma independent_substdbasis: "d \<subseteq> Basis \<Longrightarrow> independent d"
   21.15 @@ -324,7 +324,7 @@
   21.16      and "setsum (\<lambda>y. if (x = y) then P x else Q y) s = setsum Q s"
   21.17      and "setsum (\<lambda>y. if (y = x) then P y else Q y) s = setsum Q s"
   21.18      and "setsum (\<lambda>y. if (x = y) then P y else Q y) s = setsum Q s"
   21.19 -  apply (rule_tac [!] setsum_cong2)
   21.20 +  apply (rule_tac [!] setsum.cong)
   21.21    using assms
   21.22    apply auto
   21.23    done
   21.24 @@ -337,11 +337,11 @@
   21.25    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)"
   21.26      by auto
   21.27    show ?thesis
   21.28 -    unfolding * using setsum_delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
   21.29 +    unfolding * using setsum.delta[OF assms, of y "\<lambda>x. f x *\<^sub>R x"] by auto
   21.30  qed
   21.31  
   21.32  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)"
   21.33 -  by auto
   21.34 +  by (fact if_distrib)
   21.35  
   21.36  lemma dist_triangle_eq:
   21.37    fixes x y z :: "'a::real_inner"
   21.38 @@ -578,8 +578,8 @@
   21.39          setsum ua s = 1 \<and> (\<Sum>v\<in>s. ua v *\<^sub>R v) = u *\<^sub>R x + v *\<^sub>R y"
   21.40        apply (rule_tac x="sx \<union> sy" in exI)
   21.41        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)
   21.42 -      unfolding scaleR_left_distrib setsum_addf if_smult scaleR_zero_left
   21.43 -        ** setsum_restrict_set[OF xy, symmetric]
   21.44 +      unfolding scaleR_left_distrib setsum.distrib if_smult scaleR_zero_left
   21.45 +        ** setsum.inter_restrict[OF xy, symmetric]
   21.46        unfolding scaleR_scaleR[symmetric] Real_Vector_Spaces.scaleR_right.setsum [symmetric]
   21.47          and setsum_right_distrib[symmetric]
   21.48        unfolding x y
   21.49 @@ -616,7 +616,7 @@
   21.50    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"
   21.51    then show "\<exists>u. setsum u s = 1 \<and> (\<Sum>v\<in>s. u v *\<^sub>R v) = x"
   21.52      apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
   21.53 -    unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms, symmetric] and *
   21.54 +    unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms, symmetric] and *
   21.55      apply auto
   21.56      done
   21.57  qed
   21.58 @@ -673,9 +673,9 @@
   21.59          apply (rule_tac x="\<lambda>x. (if x=a then v else 0) + u x" in exI)
   21.60          unfolding setsum_clauses(2)[OF fin]
   21.61          apply simp
   21.62 -        unfolding scaleR_left_distrib and setsum_addf
   21.63 +        unfolding scaleR_left_distrib and setsum.distrib
   21.64          unfolding vu and * and scaleR_zero_left
   21.65 -        apply (auto simp add: setsum_delta[OF fin])
   21.66 +        apply (auto simp add: setsum.delta[OF fin])
   21.67          done
   21.68      next
   21.69        case False
   21.70 @@ -685,8 +685,8 @@
   21.71        from False show ?thesis
   21.72          apply (rule_tac x="\<lambda>x. if x=a then v else u x" in exI)
   21.73          unfolding setsum_clauses(2)[OF fin] and * using vu
   21.74 -        using setsum_cong2[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)]
   21.75 -        using setsum_cong2[of s u "\<lambda>x. if x = a then v else u x", OF **(1)]
   21.76 +        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)]
   21.77 +        using setsum.cong [of s _ u "\<lambda>x. if x = a then v else u x", OF _ **(1)]
   21.78          apply auto
   21.79          done
   21.80      qed
   21.81 @@ -777,7 +777,7 @@
   21.82      apply (rule conjI) using as(1) apply simp
   21.83      apply (erule conjI)
   21.84      using as(1)
   21.85 -    apply (simp add: setsum_reindex[unfolded inj_on_def] scaleR_right_diff_distrib
   21.86 +    apply (simp add: setsum.reindex[unfolded inj_on_def] scaleR_right_diff_distrib
   21.87        setsum_subtractf scaleR_left.setsum[symmetric] setsum_diff1 scaleR_left_diff_distrib)
   21.88      unfolding as
   21.89      apply simp
   21.90 @@ -797,7 +797,7 @@
   21.91      unfolding span_explicit by auto
   21.92    def f \<equiv> "(\<lambda>x. x + a) ` t"
   21.93    have f: "finite f" "f \<subseteq> s" "(\<Sum>v\<in>f. u (v - a) *\<^sub>R (v - a)) = y - a"
   21.94 -    unfolding f_def using obt by (auto simp add: setsum_reindex[unfolded inj_on_def])
   21.95 +    unfolding f_def using obt by (auto simp add: setsum.reindex[unfolded inj_on_def])
   21.96    have *: "f \<inter> {a} = {}" "f \<inter> - {a} = f"
   21.97      using f(2) assms by auto
   21.98    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"
   21.99 @@ -805,7 +805,7 @@
  21.100      apply (rule_tac x = "\<lambda>x. if x=a then 1 - setsum (\<lambda>x. u (x - a)) f else u (x - a)" in exI)
  21.101      using assms and f
  21.102      unfolding setsum_clauses(2)[OF f(1)] and if_smult
  21.103 -    unfolding setsum_cases[OF f(1), of "\<lambda>x. x = a"]
  21.104 +    unfolding setsum.If_cases[OF f(1), of "\<lambda>x. x = a"]
  21.105      apply (auto simp add: setsum_subtractf scaleR_left.setsum algebra_simps *)
  21.106      done
  21.107  qed
  21.108 @@ -1284,7 +1284,7 @@
  21.109      unfolding affine_dependent_explicit by auto
  21.110    then show ?rhs
  21.111      apply (rule_tac x="\<lambda>x. if x\<in>t then u x else 0" in exI)
  21.112 -    apply auto unfolding * and setsum_restrict_set[OF assms, symmetric]
  21.113 +    apply auto unfolding * and setsum.inter_restrict[OF assms, symmetric]
  21.114      unfolding Int_absorb1[OF `t\<subseteq>s`]
  21.115      apply auto
  21.116      done
  21.117 @@ -1708,8 +1708,8 @@
  21.118      apply (rule, rule)
  21.119      defer
  21.120      apply rule
  21.121 -    unfolding * and setsum_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
  21.122 -      setsum_reindex[OF inj] and o_def Collect_mem_eq
  21.123 +    unfolding * and setsum.If_cases[OF finite_atLeastAtMost[of 1 "k1 + k2"]] and
  21.124 +      setsum.reindex[OF inj] and o_def Collect_mem_eq
  21.125      unfolding scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] setsum_right_distrib[symmetric]
  21.126    proof -
  21.127      fix i
  21.128 @@ -1744,7 +1744,7 @@
  21.129    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"
  21.130      apply (rule_tac x="\<lambda>y. if x=y then 1 else 0" in exI)
  21.131      apply auto
  21.132 -    unfolding setsum_delta'[OF assms] and setsum_delta''[OF assms]
  21.133 +    unfolding setsum.delta'[OF assms] and setsum_delta''[OF assms]
  21.134      apply auto
  21.135      done
  21.136  next
  21.137 @@ -1761,11 +1761,11 @@
  21.138    }
  21.139    moreover
  21.140    have "(\<Sum>x\<in>s. u * ux x + v * uy x) = 1"
  21.141 -    unfolding setsum_addf and setsum_right_distrib[symmetric] and ux(2) uy(2)
  21.142 +    unfolding setsum.distrib and setsum_right_distrib[symmetric] and ux(2) uy(2)
  21.143      using uv(3) by auto
  21.144    moreover
  21.145    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)"
  21.146 -    unfolding scaleR_left_distrib and setsum_addf and scaleR_scaleR[symmetric]
  21.147 +    unfolding scaleR_left_distrib and setsum.distrib and scaleR_scaleR[symmetric]
  21.148        and scaleR_right.setsum [symmetric]
  21.149      by auto
  21.150    ultimately
  21.151 @@ -1877,8 +1877,8 @@
  21.152        unfolding setsum_image_gen[OF *(1), of "\<lambda>x. u (f x) *\<^sub>R f x" f]
  21.153          and setsum_image_gen[OF *(1), of "\<lambda>x. u (f x)" f]
  21.154        unfolding f
  21.155 -      using setsum_cong2[of 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"]
  21.156 -      using setsum_cong2 [of s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
  21.157 +      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"]
  21.158 +      using setsum.cong [of s s "\<lambda>y. (\<Sum>x\<in>{x \<in> {1..card s}. f x = y}. u (f x))" u]
  21.159        unfolding obt(4,5)
  21.160        by auto
  21.161      ultimately
  21.162 @@ -1938,7 +1938,7 @@
  21.163      by auto
  21.164    show ?lhs
  21.165      apply (rule_tac x = "\<lambda>x. (if a = x then v else 0) + u x" in exI)
  21.166 -    unfolding scaleR_left_distrib and setsum_addf and setsum_delta''[OF fin] and setsum_delta'[OF fin]
  21.167 +    unfolding scaleR_left_distrib and setsum.distrib and setsum_delta''[OF fin] and setsum.delta'[OF fin]
  21.168      unfolding setsum_clauses(2)[OF assms]
  21.169      using uv and uv(2)[THEN bspec[where x=a]] and `a\<in>s`
  21.170      apply auto
  21.171 @@ -1953,9 +1953,9 @@
  21.172    moreover
  21.173    have "(\<Sum>x\<in>s. if a = x then v else u x) = setsum u s"
  21.174      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)"
  21.175 -    apply (rule_tac setsum_cong2)
  21.176 +    apply (rule_tac setsum.cong) apply rule
  21.177      defer
  21.178 -    apply (rule_tac setsum_cong2)
  21.179 +    apply (rule_tac setsum.cong) apply rule
  21.180      using `a \<notin> s`
  21.181      apply auto
  21.182      done
  21.183 @@ -2088,7 +2088,7 @@
  21.184    then have "finite (insert a t)" and "insert a t \<subseteq> insert a s"
  21.185      by auto
  21.186    moreover have *: "\<And>P Q. (\<Sum>x\<in>t. (if x = a then P x else Q x)) = (\<Sum>x\<in>t. Q x)"
  21.187 -    apply (rule setsum_cong2)
  21.188 +    apply (rule setsum.cong)
  21.189      using `a\<notin>s` `t\<subseteq>s`
  21.190      apply auto
  21.191      done
  21.192 @@ -2106,15 +2106,15 @@
  21.193      apply auto
  21.194      done
  21.195    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)"
  21.196 -    apply (rule setsum_cong2)
  21.197 +    apply (rule setsum.cong)
  21.198      using `a\<notin>s` `t\<subseteq>s`
  21.199      apply auto
  21.200      done
  21.201    have "(\<Sum>x\<in>t. u (x - a)) *\<^sub>R a = (\<Sum>v\<in>t. u (v - a) *\<^sub>R v)"
  21.202      unfolding scaleR_left.setsum
  21.203 -    unfolding t_def and setsum_reindex[OF inj] and o_def
  21.204 +    unfolding t_def and setsum.reindex[OF inj] and o_def
  21.205      using obt(5)
  21.206 -    by (auto simp add: setsum_addf scaleR_right_distrib)
  21.207 +    by (auto simp add: setsum.distrib scaleR_right_distrib)
  21.208    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"
  21.209      unfolding setsum_clauses(2)[OF fin]
  21.210      using `a\<notin>s` `t\<subseteq>s`
  21.211 @@ -2250,7 +2250,7 @@
  21.212          apply auto
  21.213          done
  21.214        then have "setsum w s > 0"
  21.215 -        unfolding setsum_diff1'[OF obt(1) `v\<in>s`]
  21.216 +        unfolding setsum.remove[OF obt(1) `v\<in>s`]
  21.217          using as[THEN bspec[where x=v]] and `v\<in>s`
  21.218          using `w v \<noteq> 0`
  21.219          by auto
  21.220 @@ -2293,11 +2293,11 @@
  21.221        using Min_in[OF _ `i\<noteq>{}`] and obt(1) unfolding i_def t_def by auto
  21.222      then have a: "a \<in> s" "u a + t * w a = 0" by auto
  21.223      have *: "\<And>f. setsum f (s - {a}) = setsum f s - ((f a)::'b::ab_group_add)"
  21.224 -      unfolding setsum_diff1'[OF obt(1) `a\<in>s`] by auto
  21.225 +      unfolding setsum.remove[OF obt(1) `a\<in>s`] by auto
  21.226      have "(\<Sum>v\<in>s. u v + t * w v) = 1"
  21.227 -      unfolding setsum_addf wv(1) setsum_right_distrib[symmetric] obt(5) by auto
  21.228 +      unfolding setsum.distrib wv(1) setsum_right_distrib[symmetric] obt(5) by auto
  21.229      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"
  21.230 -      unfolding setsum_addf obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
  21.231 +      unfolding setsum.distrib obt(6) scaleR_scaleR[symmetric] scaleR_right.setsum [symmetric] wv(4)
  21.232        using a(2) [THEN eq_neg_iff_add_eq_0 [THEN iffD2]] by simp
  21.233      ultimately have "?P (n - 1)"
  21.234        apply (rule_tac x="(s - {a})" in exI)
  21.235 @@ -3762,10 +3762,10 @@
  21.236      have *: "inj_on (\<lambda>v. v + (y - a)) t"
  21.237        unfolding inj_on_def by auto
  21.238      have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a))) = 1"
  21.239 -      unfolding setsum_reindex[OF *] o_def using obt(4) by auto
  21.240 +      unfolding setsum.reindex[OF *] o_def using obt(4) by auto
  21.241      moreover have "(\<Sum>v\<in>(\<lambda>v. v + (y - a)) ` t. u (v - (y - a)) *\<^sub>R v) = y"
  21.242 -      unfolding setsum_reindex[OF *] o_def using obt(4,5)
  21.243 -      by (simp add: setsum_addf setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib)
  21.244 +      unfolding setsum.reindex[OF *] o_def using obt(4,5)
  21.245 +      by (simp add: setsum.distrib setsum_subtractf scaleR_left.setsum[symmetric] scaleR_right_distrib)
  21.246      ultimately
  21.247      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"
  21.248        apply (rule_tac x="(\<lambda>v. v + (y - a)) ` t" in exI)
  21.249 @@ -4776,7 +4776,7 @@
  21.250      by blast
  21.251    then show ?thesis
  21.252      apply (rule_tac x="\<lambda>v. if v\<in>s then u v else 0" in exI)
  21.253 -    unfolding if_smult scaleR_zero_left and setsum_restrict_set[OF assms(1), symmetric]
  21.254 +    unfolding if_smult scaleR_zero_left and setsum.inter_restrict[OF assms(1), symmetric]
  21.255      apply (auto simp add: Int_absorb1)
  21.256      done
  21.257  qed
  21.258 @@ -4789,8 +4789,8 @@
  21.259    have *: "\<And>x. (if f x < 0 then f x else 0) + (if 0 < f x then f x else 0) = f x"
  21.260      by auto
  21.261    show ?thesis
  21.262 -    unfolding real_add_eq_0_iff[symmetric] and setsum_restrict_set''[OF assms(1)]
  21.263 -      and setsum_addf[symmetric] and *
  21.264 +    unfolding real_add_eq_0_iff[symmetric] and setsum.inter_filter[OF assms(1)]
  21.265 +      and setsum.distrib[symmetric] and *
  21.266      using assms(2)
  21.267      apply assumption
  21.268      done
  21.269 @@ -4805,8 +4805,8 @@
  21.270    have *: "\<And>x. (if 0 < g x then f x else 0) + (if g x < 0 then f x else 0) = f x"
  21.271      using assms(3) by auto
  21.272    show ?thesis
  21.273 -    unfolding eq_neg_iff_add_eq_0 and setsum_restrict_set''[OF assms(1)]
  21.274 -      and setsum_addf[symmetric] and *
  21.275 +    unfolding eq_neg_iff_add_eq_0 and setsum.inter_filter[OF assms(1)]
  21.276 +      and setsum.distrib[symmetric] and *
  21.277      using assms(2)
  21.278      apply assumption
  21.279      done
  21.280 @@ -4837,7 +4837,7 @@
  21.281          apply auto
  21.282          done
  21.283        then show ?thesis
  21.284 -        unfolding setsum_delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto
  21.285 +        unfolding setsum.delta[OF assms(1)] using uv(2) and `u v < 0` and uv(1) by auto
  21.286      qed
  21.287    qed (insert setsum_nonneg_eq_0_iff[of _ u, OF fin(1)] uv(2-3), auto)
  21.288  
  21.289 @@ -4850,14 +4850,14 @@
  21.290    moreover have "setsum u ({x \<in> c. 0 < u x} \<union> {x \<in> c. u x < 0}) = setsum u c"
  21.291      "(\<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)"
  21.292      using assms(1)
  21.293 -    apply (rule_tac[!] setsum_mono_zero_left)
  21.294 +    apply (rule_tac[!] setsum.mono_neutral_left)
  21.295      apply auto
  21.296      done
  21.297    then have "setsum u {x \<in> c. 0 < u x} = - setsum u {x \<in> c. 0 > u x}"
  21.298      "(\<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)"
  21.299      unfolding eq_neg_iff_add_eq_0
  21.300      using uv(1,4)
  21.301 -    by (auto simp add: setsum_Un_zero[OF fin, symmetric])
  21.302 +    by (auto simp add: setsum.union_inter_neutral[OF fin, symmetric])
  21.303    moreover have "\<forall>x\<in>{v \<in> c. u v < 0}. 0 \<le> inverse (setsum u {x \<in> c. 0 < u x}) * - u x"
  21.304      apply rule
  21.305      apply (rule mult_nonneg_nonneg)
  21.306 @@ -5773,7 +5773,7 @@
  21.307  qed
  21.308  
  21.309  lemma inner_setsum_Basis[simp]: "i \<in> Basis \<Longrightarrow> (\<Sum>Basis) \<bullet> i = 1"
  21.310 -  by (simp add: inner_setsum_left setsum_cases inner_Basis)
  21.311 +  by (simp add: inner_setsum_left setsum.If_cases inner_Basis)
  21.312  
  21.313  lemma convex_set_plus:
  21.314    assumes "convex s" and "convex t" shows "convex (s + t)"
  21.315 @@ -5806,7 +5806,7 @@
  21.316    apply (rule_tac x="fun_upd f x a" in exI)
  21.317    apply simp
  21.318    apply (rule_tac f="\<lambda>x. a + x" in arg_cong)
  21.319 -  apply (rule setsum_cong [OF refl])
  21.320 +  apply (rule setsum.cong [OF refl])
  21.321    apply clarsimp
  21.322    apply (fast intro: set_plus_intro)
  21.323    done
  21.324 @@ -5820,10 +5820,10 @@
  21.325    apply (subgoal_tac "(\<Sum>x\<in>Basis. f x \<bullet> i) = f i \<bullet> i")
  21.326    apply (drule (1) bspec)
  21.327    apply clarsimp
  21.328 -  apply (frule setsum_diff1' [OF finite_Basis])
  21.329 +  apply (frule setsum.remove [OF finite_Basis])
  21.330    apply (erule trans)
  21.331    apply simp
  21.332 -  apply (rule setsum_0')
  21.333 +  apply (rule setsum.neutral)
  21.334    apply clarsimp
  21.335    apply (frule_tac x=i in bspec, assumption)
  21.336    apply (drule_tac x=x in bspec, assumption)
  21.337 @@ -5860,7 +5860,7 @@
  21.338    (is "?int = convex hull ?points")
  21.339  proof -
  21.340    have One[simp]: "\<And>i. i \<in> Basis \<Longrightarrow> One \<bullet> i = 1"
  21.341 -    by (simp add: One_def inner_setsum_left setsum_cases inner_Basis)
  21.342 +    by (simp add: One_def inner_setsum_left setsum.If_cases inner_Basis)
  21.343    have "?int = {x. \<forall>i\<in>Basis. x \<bullet> i \<in> cbox 0 1}"
  21.344      by (auto simp: cbox_def)
  21.345    also have "\<dots> = (\<Sum>i\<in>Basis. (\<lambda>x. x *\<^sub>R i) ` cbox 0 1)"
  21.346 @@ -6154,7 +6154,7 @@
  21.347        unfolding c_def convex_hull_set_setsum
  21.348        apply (subst convex_hull_linear_image [symmetric])
  21.349        apply (simp add: linear_iff scaleR_add_left)
  21.350 -      apply (rule setsum_cong [OF refl])
  21.351 +      apply (rule setsum.cong [OF refl])
  21.352        apply (rule image_cong [OF _ refl])
  21.353        apply (rule convex_hull_eq_real_cbox)
  21.354        apply (cut_tac `0 < d`, simp)
  21.355 @@ -6625,7 +6625,7 @@
  21.356        by auto
  21.357      then have **: "setsum u ?D = setsum (op \<bullet> x) ?D"
  21.358        apply -
  21.359 -      apply (rule setsum_cong2)
  21.360 +      apply (rule setsum.cong)
  21.361        using assms
  21.362        apply auto
  21.363        done
  21.364 @@ -6694,13 +6694,13 @@
  21.365        by (auto simp: SOME_Basis inner_Basis inner_simps)
  21.366      then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis =
  21.367        setsum (\<lambda>i. x\<bullet>i + (if (SOME i. i\<in>Basis) = i then e/2 else 0)) Basis"
  21.368 -      apply (rule_tac setsum_cong)
  21.369 +      apply (rule_tac setsum.cong)
  21.370        apply auto
  21.371        done
  21.372      have "setsum (op \<bullet> x) Basis < setsum (op \<bullet> (x + (e / 2) *\<^sub>R (SOME i. i\<in>Basis))) Basis"
  21.373 -      unfolding * setsum_addf
  21.374 +      unfolding * setsum.distrib
  21.375        using `e > 0` DIM_positive[where 'a='a]
  21.376 -      apply (subst setsum_delta')
  21.377 +      apply (subst setsum.delta')
  21.378        apply (auto simp: SOME_Basis)
  21.379        done
  21.380      also have "\<dots> \<le> 1"
  21.381 @@ -6744,7 +6744,7 @@
  21.382        then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
  21.383      qed
  21.384      also have "\<dots> \<le> 1"
  21.385 -      unfolding setsum_addf setsum_constant real_eq_of_nat
  21.386 +      unfolding setsum.distrib setsum_constant real_eq_of_nat
  21.387        by (auto simp add: Suc_le_eq)
  21.388      finally show "(\<forall>i\<in>Basis. 0 \<le> y \<bullet> i) \<and> setsum (op \<bullet> y) Basis \<le> 1"
  21.389      proof safe
  21.390 @@ -6774,7 +6774,7 @@
  21.391      assume i: "i \<in> Basis"
  21.392      have "?a \<bullet> i = inverse (2 * real DIM('a))"
  21.393        by (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"])
  21.394 -         (simp_all add: setsum_cases i) }
  21.395 +         (simp_all add: setsum.If_cases i) }
  21.396    note ** = this
  21.397    show ?thesis
  21.398      apply (rule that[of ?a])
  21.399 @@ -6786,7 +6786,8 @@
  21.400        unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive)
  21.401    next
  21.402      have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real DIM('a))) ?D"
  21.403 -      apply (rule setsum_cong2, rule **)
  21.404 +      apply (rule setsum.cong)
  21.405 +      apply rule
  21.406        apply auto
  21.407        done
  21.408      also have "\<dots> < 1"
  21.409 @@ -6859,16 +6860,16 @@
  21.410            using a d by (auto simp: inner_simps inner_Basis)
  21.411          then have *: "setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d =
  21.412            setsum (\<lambda>i. x\<bullet>i + (if a = i then e/2 else 0)) d"
  21.413 -          using d by (intro setsum_cong) auto
  21.414 +          using d by (intro setsum.cong) auto
  21.415          have "a \<in> Basis"
  21.416            using `a \<in> d` d by auto
  21.417          then have h1: "(\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> (x + (e / 2) *\<^sub>R a) \<bullet> i = 0)"
  21.418            using x0 d `a\<in>d` by (auto simp add: inner_add_left inner_Basis)
  21.419          have "setsum (op \<bullet> x) d < setsum (op \<bullet> (x + (e / 2) *\<^sub>R a)) d"
  21.420 -          unfolding * setsum_addf
  21.421 +          unfolding * setsum.distrib
  21.422            using `e > 0` `a \<in> d`
  21.423            using `finite d`
  21.424 -          by (auto simp add: setsum_delta')
  21.425 +          by (auto simp add: setsum.delta')
  21.426          also have "\<dots> \<le> 1"
  21.427            using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"]
  21.428            by auto
  21.429 @@ -6927,7 +6928,7 @@
  21.430            then show "y \<bullet> i \<le> x \<bullet> i + ?d" by auto
  21.431          qed
  21.432          also have "\<dots> \<le> 1"
  21.433 -          unfolding setsum_addf setsum_constant real_eq_of_nat
  21.434 +          unfolding setsum.distrib setsum_constant real_eq_of_nat
  21.435            using `0 < card d`
  21.436            by auto
  21.437          finally show "setsum (op \<bullet> y) d \<le> 1" .
  21.438 @@ -6976,10 +6977,10 @@
  21.439      have "?a \<bullet> i = inverse (2 * real (card d))"
  21.440        apply (rule trans[of _ "setsum (\<lambda>j. if i = j then inverse (2 * real (card d)) else 0) ?D"])
  21.441        unfolding inner_setsum_left
  21.442 -      apply (rule setsum_cong2)
  21.443 -      using `i \<in> d` `finite d` setsum_delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
  21.444 +      apply (rule setsum.cong)
  21.445 +      using `i \<in> d` `finite d` setsum.delta'[of d i "(\<lambda>k. inverse (2 * real (card d)))"]
  21.446          d1 assms(2)
  21.447 -      by (auto simp: inner_simps inner_Basis set_rev_mp[OF _ assms(2)])
  21.448 +      by (auto simp: inner_Basis set_rev_mp[OF _ assms(2)])
  21.449    }
  21.450    note ** = this
  21.451    show ?thesis
  21.452 @@ -6995,7 +6996,7 @@
  21.453      finally show "0 < ?a \<bullet> i" by auto
  21.454    next
  21.455      have "setsum (op \<bullet> ?a) ?D = setsum (\<lambda>i. inverse (2 * real (card d))) ?D"
  21.456 -      by (rule setsum_cong2) (rule **)
  21.457 +      by (rule setsum.cong) (rule refl, rule **)
  21.458      also have "\<dots> < 1"
  21.459        unfolding setsum_constant real_eq_of_nat divide_real_def[symmetric]
  21.460        by (auto simp add: field_simps)
  21.461 @@ -8395,13 +8396,13 @@
  21.462        assume "x \<in> S i"
  21.463        def c \<equiv> "\<lambda>j. if j = i then 1::real else 0"
  21.464        then have *: "setsum c I = 1"
  21.465 -        using `finite I` `i \<in> I` setsum_delta[of I i "\<lambda>j::'a. 1::real"]
  21.466 +        using `finite I` `i \<in> I` setsum.delta[of I i "\<lambda>j::'a. 1::real"]
  21.467          by auto
  21.468        def s \<equiv> "\<lambda>j. if j = i then x else p j"
  21.469        then have "\<forall>j. c j *\<^sub>R s j = (if j = i then x else 0)"
  21.470          using c_def by (auto simp add: algebra_simps)
  21.471        then have "x = setsum (\<lambda>i. c i *\<^sub>R s i) I"
  21.472 -        using s_def c_def `finite I` `i \<in> I` setsum_delta[of I i "\<lambda>j::'a. x"]
  21.473 +        using s_def c_def `finite I` `i \<in> I` setsum.delta[of I i "\<lambda>j::'a. x"]
  21.474          by auto
  21.475        then have "x \<in> ?rhs"
  21.476          apply auto
  21.477 @@ -8434,7 +8435,7 @@
  21.478      moreover have "setsum (\<lambda>i. v * d i) I = v * setsum d I"
  21.479        by (simp add: setsum_right_distrib)
  21.480      ultimately have sum1: "setsum e I = 1"
  21.481 -      using e_def xc yc uv by (simp add: setsum_addf)
  21.482 +      using e_def xc yc uv by (simp add: setsum.distrib)
  21.483      def q \<equiv> "\<lambda>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"
  21.484      {
  21.485        fix i
  21.486 @@ -8477,7 +8478,7 @@
  21.487      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"
  21.488        by auto
  21.489      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"
  21.490 -      using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum_addf)
  21.491 +      using xc yc by (simp add: algebra_simps scaleR_right.setsum setsum.distrib)
  21.492      also have "\<dots> = setsum (\<lambda>i. e i *\<^sub>R q i) I"
  21.493        using * by auto
  21.494      finally have "u *\<^sub>R x + v *\<^sub>R y = setsum (\<lambda>i. (e i) *\<^sub>R (q i)) I"
    22.1 --- a/src/HOL/Multivariate_Analysis/Determinants.thy	Fri Jun 27 22:08:55 2014 +0200
    22.2 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Sat Jun 28 09:16:42 2014 +0200
    22.3 @@ -12,13 +12,6 @@
    22.4  
    22.5  subsection{* First some facts about products*}
    22.6  
    22.7 -lemma setprod_insert_eq:
    22.8 -  "finite A \<Longrightarrow> setprod f (insert a A) = (if a \<in> A then setprod f A else f a * setprod f A)"
    22.9 -  apply clarsimp
   22.10 -  apply (subgoal_tac "insert a A = A")
   22.11 -  apply auto
   22.12 -  done
   22.13 -
   22.14  lemma setprod_add_split:
   22.15    fixes m n :: nat
   22.16    assumes mn: "m \<le> n + 1"
   22.17 @@ -33,7 +26,7 @@
   22.18      by auto
   22.19    have f: "finite ?B" "finite ?C"
   22.20      by simp_all
   22.21 -  from setprod_Un_disjoint[OF f dj, of f, unfolded un] show ?thesis .
   22.22 +  from setprod.union_disjoint[OF f dj, of f, unfolded un] show ?thesis .
   22.23  qed
   22.24  
   22.25  
   22.26 @@ -82,7 +75,7 @@
   22.27    assumes fS: "finite S"
   22.28      and f: "\<forall>x\<in>S. f x \<ge> 0 \<and> f x \<le> 1"
   22.29    shows "setprod f S \<le> 1"
   22.30 -  using setprod_le[OF fS f] unfolding setprod_1 .
   22.31 +  using setprod_le[OF fS f] unfolding setprod.neutral_const .
   22.32  
   22.33  
   22.34  subsection {* Trace *}
   22.35 @@ -97,14 +90,14 @@
   22.36    by (simp add: trace_def mat_def)
   22.37  
   22.38  lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
   22.39 -  by (simp add: trace_def setsum_addf)
   22.40 +  by (simp add: trace_def setsum.distrib)
   22.41  
   22.42  lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
   22.43    by (simp add: trace_def setsum_subtractf)
   22.44  
   22.45  lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)"
   22.46    apply (simp add: trace_def matrix_matrix_mult_def)
   22.47 -  apply (subst setsum_commute)
   22.48 +  apply (subst setsum.commute)
   22.49    apply (simp add: mult_commute)
   22.50    done
   22.51  
   22.52 @@ -149,7 +142,7 @@
   22.53        setprod (\<lambda>i. ?di (transpose A) i (inv p i)) (p ` ?U)"
   22.54        by simp
   22.55      also have "\<dots> = setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U"
   22.56 -      unfolding setprod_reindex[OF pi] ..
   22.57 +      unfolding setprod.reindex[OF pi] ..
   22.58      also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U"
   22.59      proof -
   22.60        {
   22.61 @@ -161,7 +154,7 @@
   22.62        }
   22.63        then show "setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) \<circ> p) ?U =
   22.64          setprod (\<lambda>i. ?di A i (p i)) ?U"
   22.65 -        by (auto intro: setprod_cong)
   22.66 +        by (auto intro: setprod.cong)
   22.67      qed
   22.68      finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) =
   22.69        of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)"
   22.70 @@ -170,7 +163,8 @@
   22.71    then show ?thesis
   22.72      unfolding det_def
   22.73      apply (subst setsum_permutations_inverse)
   22.74 -    apply (rule setsum_cong2)
   22.75 +    apply (rule setsum.cong)
   22.76 +    apply (rule refl)
   22.77      apply blast
   22.78      done
   22.79  qed
   22.80 @@ -202,7 +196,7 @@
   22.81    }
   22.82    then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
   22.83      by blast
   22.84 -  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   22.85 +  from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
   22.86      unfolding det_def by (simp add: sign_id)
   22.87  qed
   22.88  
   22.89 @@ -233,7 +227,7 @@
   22.90    }
   22.91    then have p0: "\<forall>p \<in> ?PU -{id}. ?pp p = 0"
   22.92      by blast
   22.93 -  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   22.94 +  from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
   22.95      unfolding det_def by (simp add: sign_id)
   22.96  qed
   22.97  
   22.98 @@ -263,7 +257,7 @@
   22.99    }
  22.100    then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"
  22.101      by blast
  22.102 -  from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
  22.103 +  from setsum.mono_neutral_cong_left[OF fPU id0 p0] show ?thesis
  22.104      unfolding det_def by (simp add: sign_id)
  22.105  qed
  22.106  
  22.107 @@ -279,7 +273,7 @@
  22.108        using i by (vector mat_def)
  22.109    }
  22.110    then have th: "setprod (\<lambda>i. ?f i i) ?U = setprod (\<lambda>x. 1) ?U"
  22.111 -    by (auto intro: setprod_cong)
  22.112 +    by (auto intro: setprod.cong)
  22.113    {
  22.114      fix i j
  22.115      assume i: "i \<in> ?U" and j: "j \<in> ?U" and ij: "i \<noteq> j"
  22.116 @@ -289,7 +283,7 @@
  22.117    then have "det ?A = setprod (\<lambda>i. ?f i i) ?U"
  22.118      using det_diagonal by blast
  22.119    also have "\<dots> = 1"
  22.120 -    unfolding th setprod_1 ..
  22.121 +    unfolding th setprod.neutral_const ..
  22.122    finally show ?thesis .
  22.123  qed
  22.124  
  22.125 @@ -302,7 +296,7 @@
  22.126    shows "det (\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
  22.127    apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
  22.128    apply (subst sum_permutations_compose_right[OF p])
  22.129 -proof (rule setsum_cong2)
  22.130 +proof (rule setsum.cong)
  22.131    let ?U = "UNIV :: 'n set"
  22.132    let ?PU = "{p. p permutes ?U}"
  22.133    fix q
  22.134 @@ -325,7 +319,7 @@
  22.135    show "of_int (sign (q \<circ> p)) * setprod (\<lambda>i. A$ p i$ (q \<circ> p) i) ?U =
  22.136      of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
  22.137      by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
  22.138 -qed
  22.139 +qed rule
  22.140  
  22.141  lemma det_permute_columns:
  22.142    fixes A :: "'a::comm_ring_1^'n^'n"
  22.143 @@ -377,7 +371,7 @@
  22.144    shows "det A = 0"
  22.145    using r
  22.146    apply (simp add: row_def det_def vec_eq_iff)
  22.147 -  apply (rule setsum_0')
  22.148 +  apply (rule setsum.neutral)
  22.149    apply (auto simp: sign_nz)
  22.150    done
  22.151  
  22.152 @@ -395,8 +389,8 @@
  22.153    shows "det((\<chi> i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) =
  22.154      det((\<chi> i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) +
  22.155      det((\<chi> i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)"
  22.156 -  unfolding det_def vec_lambda_beta setsum_addf[symmetric]
  22.157 -proof (rule setsum_cong2)
  22.158 +  unfolding det_def vec_lambda_beta setsum.distrib[symmetric]
  22.159 +proof (rule setsum.cong)
  22.160    let ?U = "UNIV :: 'n set"
  22.161    let ?pU = "{p. p permutes ?U}"
  22.162    let ?f = "(\<lambda>i. if i = k then a i + b i else c i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
  22.163 @@ -418,14 +412,14 @@
  22.164    then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
  22.165      and th2: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?h i $ p i) ?Uk"
  22.166      apply -
  22.167 -    apply (rule setprod_cong, simp_all)+
  22.168 +    apply (rule setprod.cong, simp_all)+
  22.169      done
  22.170    have th3: "finite ?Uk" "k \<notin> ?Uk"
  22.171      by auto
  22.172    have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
  22.173      unfolding kU[symmetric] ..
  22.174    also have "\<dots> = ?f k $ p k * setprod (\<lambda>i. ?f i $ p i) ?Uk"
  22.175 -    apply (rule setprod_insert)
  22.176 +    apply (rule setprod.insert)
  22.177      apply simp
  22.178      apply blast
  22.179      done
  22.180 @@ -434,20 +428,20 @@
  22.181    also have "\<dots> = (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\<lambda>i. ?h i $ p i) ?Uk)"
  22.182      by (metis th1 th2)
  22.183    also have "\<dots> = setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk) + setprod (\<lambda>i. ?h i $ p i) (insert k ?Uk)"
  22.184 -    unfolding  setprod_insert[OF th3] by simp
  22.185 +    unfolding  setprod.insert[OF th3] by simp
  22.186    finally have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?g i $ p i) ?U + setprod (\<lambda>i. ?h i $ p i) ?U"
  22.187      unfolding kU[symmetric] .
  22.188    then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U =
  22.189      of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U + of_int (sign p) * setprod (\<lambda>i. ?h i $ p i) ?U"
  22.190      by (simp add: field_simps)
  22.191 -qed
  22.192 +qed rule
  22.193  
  22.194  lemma det_row_mul:
  22.195    fixes a b :: "'n::finite \<Rightarrow> _ ^ 'n"
  22.196    shows "det((\<chi> i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) =
  22.197      c * det((\<chi> i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)"
  22.198    unfolding det_def vec_lambda_beta setsum_right_distrib
  22.199 -proof (rule setsum_cong2)
  22.200 +proof (rule setsum.cong)
  22.201    let ?U = "UNIV :: 'n set"
  22.202    let ?pU = "{p. p permutes ?U}"
  22.203    let ?f = "(\<lambda>i. if i = k then c*s a i else b i)::'n \<Rightarrow> 'a::comm_ring_1^'n"
  22.204 @@ -467,7 +461,7 @@
  22.205    }
  22.206    then have th1: "setprod (\<lambda>i. ?f i $ p i) ?Uk = setprod (\<lambda>i. ?g i $ p i) ?Uk"
  22.207      apply -
  22.208 -    apply (rule setprod_cong)
  22.209 +    apply (rule setprod.cong)
  22.210      apply simp_all
  22.211      done
  22.212    have th3: "finite ?Uk" "k \<notin> ?Uk"
  22.213 @@ -475,7 +469,7 @@
  22.214    have "setprod (\<lambda>i. ?f i $ p i) ?U = setprod (\<lambda>i. ?f i $ p i) (insert k ?Uk)"
  22.215      unfolding kU[symmetric] ..
  22.216    also have "\<dots> = ?f k $ p k  * setprod (\<lambda>i. ?f i $ p i) ?Uk"
  22.217 -    apply (rule setprod_insert)
  22.218 +    apply (rule setprod.insert)
  22.219      apply simp
  22.220      apply blast
  22.221      done
  22.222 @@ -484,13 +478,13 @@
  22.223    also have "\<dots> = c* (a k $ p k * setprod (\<lambda>i. ?g i $ p i) ?Uk)"
  22.224      unfolding th1 by (simp add: mult_ac)
  22.225    also have "\<dots> = c* (setprod (\<lambda>i. ?g i $ p i) (insert k ?Uk))"
  22.226 -    unfolding setprod_insert[OF th3] by simp
  22.227 +    unfolding setprod.insert[OF th3] by simp
  22.228    finally have "setprod (\<lambda>i. ?f i $ p i) ?U = c* (setprod (\<lambda>i. ?g i $ p i) ?U)"
  22.229      unfolding kU[symmetric] .
  22.230    then show "of_int (sign p) * setprod (\<lambda>i. ?f i $ p i) ?U =
  22.231      c * (of_int (sign p) * setprod (\<lambda>i. ?g i $ p i) ?U)"
  22.232      by (simp add: field_simps)
  22.233 -qed
  22.234 +qed rule
  22.235  
  22.236  lemma det_row_0:
  22.237    fixes b :: "'n::finite \<Rightarrow> _ ^ 'n"
  22.238 @@ -617,7 +611,7 @@
  22.239    case 1
  22.240    then show ?case
  22.241      apply simp
  22.242 -    unfolding setsum_empty det_row_0[of k]
  22.243 +    unfolding setsum.empty det_row_0[of k]
  22.244      apply rule
  22.245      done
  22.246  next
  22.247 @@ -695,7 +689,7 @@
  22.248       (\<Sum>(j, f)\<in>S \<times> ?F T. det (\<chi> i. if i \<in> T then a i (f i)
  22.249                                  else if i = z then a i j
  22.250                                  else c i))"
  22.251 -    unfolding insert.hyps unfolding setsum_cartesian_product by blast
  22.252 +    unfolding insert.hyps unfolding setsum.cartesian_product by blast
  22.253    show ?case unfolding tha
  22.254      using `z \<notin> T`
  22.255      by (intro setsum.reindex_bij_witness[where i="?k" and j="?h"])
  22.256 @@ -722,7 +716,7 @@
  22.257  lemma det_rows_mul:
  22.258    "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
  22.259      setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
  22.260 -proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
  22.261 +proof (simp add: det_def setsum_right_distrib cong add: setprod.cong, rule setsum.cong)
  22.262    let ?U = "UNIV :: 'n set"
  22.263    let ?PU = "{p. p permutes ?U}"
  22.264    fix p
  22.265 @@ -731,11 +725,11 @@
  22.266    from pU have p: "p permutes ?U"
  22.267      by blast
  22.268    have "setprod (\<lambda>i. c i * a i $ p i) ?U = setprod c ?U * setprod (\<lambda>i. a i $ p i) ?U"
  22.269 -    unfolding setprod_timesf ..
  22.270 +    unfolding setprod.distrib ..
  22.271    then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
  22.272      setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))"
  22.273      by (simp add: field_simps)
  22.274 -qed
  22.275 +qed rule
  22.276  
  22.277  lemma det_mul:
  22.278    fixes A B :: "'a::linordered_idom^'n^'n"
  22.279 @@ -816,7 +810,7 @@
  22.280          (\<Prod>i\<in> ?U. (\<chi> i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) =
  22.281        (setsum (\<lambda>q. ?s p * (\<Prod>i\<in> ?U. A $ i $ p i) * (?s q * (\<Prod>i\<in> ?U. B $ i $ q i))) ?PU)"
  22.282        unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f]
  22.283 -    proof (rule setsum_cong2)
  22.284 +    proof (rule setsum.cong)
  22.285        fix q
  22.286        assume qU: "q \<in> ?PU"
  22.287        then have q: "q permutes ?U"
  22.288 @@ -835,8 +829,8 @@
  22.289          by (rule setprod_permute[OF p])
  22.290        have thp: "setprod (\<lambda>i. (\<chi> i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U =
  22.291          setprod (\<lambda>i. A$i$p i) ?U * setprod (\<lambda>i. B$i$ q (inv p i)) ?U"
  22.292 -        unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p]
  22.293 -        apply (rule setprod_cong[OF refl])
  22.294 +        unfolding th001 setprod.distrib[symmetric] o_def permutes_inverses[OF p]
  22.295 +        apply (rule setprod.cong[OF refl])
  22.296          using permutes_in_image[OF q]
  22.297          apply vector
  22.298          done
  22.299 @@ -844,16 +838,16 @@
  22.300          ?s p * (setprod (\<lambda>i. A$i$p i) ?U) * (?s (q \<circ> inv p) * setprod (\<lambda>i. B$i$(q \<circ> inv p) i) ?U)"
  22.301          using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp]
  22.302          by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose)
  22.303 -    qed
  22.304 +    qed rule
  22.305    }
  22.306    then have th2: "setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU = det A * det B"
  22.307      unfolding det_def setsum_product
  22.308 -    by (rule setsum_cong2)
  22.309 +    by (rule setsum.cong [OF refl])
  22.310    have "det (A**B) = setsum (\<lambda>f.  det (\<chi> i. A $ i $ f i *s B $ f i)) ?F"
  22.311      unfolding matrix_mul_setsum_alt det_linear_rows_setsum[OF fU]
  22.312      by simp
  22.313    also have "\<dots> = setsum (\<lambda>f. det (\<chi> i. A$i$f i *s B$f i)) ?PU"
  22.314 -    using setsum_mono_zero_cong_left[OF fF PUF zth, symmetric]
  22.315 +    using setsum.mono_neutral_cong_left[OF fF PUF zth, symmetric]
  22.316      unfolding det_rows_mul by auto
  22.317    finally show ?thesis unfolding th2 .
  22.318  qed
  22.319 @@ -902,7 +896,7 @@
  22.320        done
  22.321      from c ci
  22.322      have thr0: "- row i A = setsum (\<lambda>j. (1/ c i) *s (c j *s row j A)) (?U - {i})"
  22.323 -      unfolding setsum_diff1'[OF fU iU] setsum_cmul
  22.324 +      unfolding setsum.remove[OF fU iU] setsum_cmul
  22.325        apply -
  22.326        apply (rule vector_mul_lcancel_imp[OF ci])
  22.327        apply (auto simp add: field_simps)
  22.328 @@ -962,7 +956,7 @@
  22.329      done
  22.330    show "?lhs = x$k * det A"
  22.331      apply (subst U)
  22.332 -    unfolding setsum_insert[OF fUk kUk]
  22.333 +    unfolding setsum.insert[OF fUk kUk]
  22.334      apply (subst th00)
  22.335      unfolding add_assoc
  22.336      apply (subst det_row_add)
  22.337 @@ -980,7 +974,7 @@
  22.338  proof -
  22.339    let ?U = "UNIV :: 'n set"
  22.340    have *: "\<And>c. setsum (\<lambda>i. c i *s row i (transpose A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
  22.341 -    by (auto simp add: row_transpose intro: setsum_cong2)
  22.342 +    by (auto simp add: row_transpose intro: setsum.cong)
  22.343    show ?thesis
  22.344      unfolding matrix_mult_vsum
  22.345      unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric]
  22.346 @@ -1073,7 +1067,7 @@
  22.347        from fd[rule_format, of "axis i 1" "axis j 1", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
  22.348        have "?A$i$j = ?m1 $ i $ j"
  22.349          by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def
  22.350 -            th0 setsum_delta[OF fU] mat_def axis_def)
  22.351 +            th0 setsum.delta[OF fU] mat_def axis_def)
  22.352      }
  22.353      then have "orthogonal_matrix ?mf"
  22.354        unfolding orthogonal_matrix
  22.355 @@ -1280,8 +1274,8 @@
  22.356  
  22.357  text {* Explicit formulas for low dimensions. *}
  22.358  
  22.359 -lemma setprod_1: "setprod f {(1::nat)..1} = f 1"
  22.360 -  by simp
  22.361 +lemma setprod_neutral_const: "setprod f {(1::nat)..1} = f 1"
  22.362 +  by (fact setprod_singleton_nat_seg)
  22.363  
  22.364  lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2"
  22.365    by (simp add: eval_nat_numeral setprod_numseg mult_commute)
    23.1 --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    23.2 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    23.3 @@ -56,7 +56,7 @@
    23.4  
    23.5  lemma (in euclidean_space) inner_setsum_left_Basis[simp]:
    23.6      "b \<in> Basis \<Longrightarrow> inner (\<Sum>i\<in>Basis. f i *\<^sub>R i) b = f b"
    23.7 -  by (simp add: inner_setsum_left inner_Basis if_distrib setsum_cases)
    23.8 +  by (simp add: inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.If_cases)
    23.9  
   23.10  lemma (in euclidean_space) euclidean_eqI:
   23.11    assumes b: "\<And>b. b \<in> Basis \<Longrightarrow> inner x b = inner y b" shows "x = y"
   23.12 @@ -166,7 +166,7 @@
   23.13      by (auto intro!: inj_onI Pair_inject)
   23.14    thus ?thesis
   23.15      unfolding Basis_prod_def
   23.16 -    by (subst setsum_Un_disjoint) (auto simp: Basis_prod_def setsum_reindex)
   23.17 +    by (subst setsum.union_disjoint) (auto simp: Basis_prod_def setsum.reindex)
   23.18  qed
   23.19  
   23.20  instance proof
    24.1 --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Fri Jun 27 22:08:55 2014 +0200
    24.2 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Sat Jun 28 09:16:42 2014 +0200
    24.3 @@ -841,7 +841,7 @@
    24.4      by auto
    24.5    from bchoice[OF this] * assms show "\<forall>i\<in>A. f i = 0"
    24.6      using setsum_nonneg_eq_0_iff[of A "\<lambda>i. real (f i)"] by auto
    24.7 -qed (rule setsum_0')
    24.8 +qed (rule setsum.neutral)
    24.9  
   24.10  lemma setsum_ereal_right_distrib:
   24.11    fixes f :: "'a \<Rightarrow> ereal"
   24.12 @@ -957,7 +957,7 @@
   24.13      and "\<And>i. 0 \<le> g i"
   24.14    shows "(\<Sum>i. f i + g i) = suminf f + suminf g"
   24.15    apply (subst (1 2 3) suminf_ereal_eq_SUP)
   24.16 -  unfolding setsum_addf
   24.17 +  unfolding setsum.distrib
   24.18    apply (intro assms ereal_add_nonneg_nonneg SUP_ereal_add_pos incseq_setsumI setsum_nonneg ballI)+
   24.19    done
   24.20  
   24.21 @@ -1237,7 +1237,7 @@
   24.22      have "(\<Sum>i<n. f (i + k)) = (\<Sum>i<n. (f \<circ> (\<lambda>i. i + k)) i)"
   24.23        by simp
   24.24      also have "\<dots> = (\<Sum>i\<in>(\<lambda>i. i + k) ` {..<n}. f i)"
   24.25 -      by (subst setsum_reindex) auto
   24.26 +      by (subst setsum.reindex) auto
   24.27      also have "\<dots> \<le> setsum f {..<n + k}"
   24.28        by (intro setsum_mono3) (auto simp: f)
   24.29      finally show "(\<Sum>i<n. f (i + k)) \<le> setsum f {..<n + k}" .
    25.1 --- a/src/HOL/Multivariate_Analysis/Fashoda.thy	Fri Jun 27 22:08:55 2014 +0200
    25.2 +++ b/src/HOL/Multivariate_Analysis/Fashoda.thy	Sat Jun 28 09:16:42 2014 +0200
    25.3 @@ -25,8 +25,8 @@
    25.4  lemma interval_bij_affine:
    25.5    "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) +
    25.6      (\<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))"
    25.7 -  by (auto simp: setsum_addf[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
    25.8 -    field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong)
    25.9 +  by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
   25.10 +    field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong)
   25.11  
   25.12  lemma continuous_interval_bij:
   25.13    fixes a b :: "'a::euclidean_space"
    26.1 --- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Fri Jun 27 22:08:55 2014 +0200
    26.2 +++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy	Sat Jun 28 09:16:42 2014 +0200
    26.3 @@ -435,7 +435,7 @@
    26.4      by (simp add: inner_commute)
    26.5    show "inner (x + y) z = inner x z + inner y z"
    26.6      unfolding inner_vec_def
    26.7 -    by (simp add: inner_add_left setsum_addf)
    26.8 +    by (simp add: inner_add_left setsum.distrib)
    26.9    show "inner (scaleR r x) y = r * inner x y"
   26.10      unfolding inner_vec_def
   26.11      by (simp add: setsum_right_distrib)
   26.12 @@ -470,17 +470,17 @@
   26.13    unfolding inner_vec_def
   26.14    apply (cases "i = j")
   26.15    apply clarsimp
   26.16 -  apply (subst setsum_diff1' [where a=j], simp_all)
   26.17 -  apply (rule setsum_0', simp add: axis_def)
   26.18 -  apply (rule setsum_0', simp add: axis_def)
   26.19 +  apply (subst setsum.remove [of _ j], simp_all)
   26.20 +  apply (rule setsum.neutral, simp add: axis_def)
   26.21 +  apply (rule setsum.neutral, simp add: axis_def)
   26.22    done
   26.23  
   26.24  lemma setsum_single:
   26.25    assumes "finite A" and "k \<in> A" and "f k = y"
   26.26    assumes "\<And>i. i \<in> A \<Longrightarrow> i \<noteq> k \<Longrightarrow> f i = 0"
   26.27    shows "(\<Sum>i\<in>A. f i) = y"
   26.28 -  apply (subst setsum_diff1' [OF assms(1,2)])
   26.29 -  apply (simp add: setsum_0' assms(3,4))
   26.30 +  apply (subst setsum.remove [OF assms(1,2)])
   26.31 +  apply (simp add: setsum.neutral assms(3,4))
   26.32    done
   26.33  
   26.34  lemma inner_axis: "inner x (axis i y) = inner (x $ i) y"
    27.1 --- a/src/HOL/Multivariate_Analysis/Integration.thy	Fri Jun 27 22:08:55 2014 +0200
    27.2 +++ b/src/HOL/Multivariate_Analysis/Integration.thy	Sat Jun 28 09:16:42 2014 +0200
    27.3 @@ -527,7 +527,7 @@
    27.4    have "0 \<in> cbox 0 (One::'a)"
    27.5      unfolding mem_box by auto
    27.6    then show ?thesis
    27.7 -     unfolding content_def interval_bounds[OF *] using setprod_1 by auto
    27.8 +     unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto
    27.9   qed
   27.10  
   27.11  lemma content_pos_le[intro]:
   27.12 @@ -1681,7 +1681,7 @@
   27.13    note assm = tagged_division_ofD[OF assms(1)]
   27.14    show ?thesis
   27.15      unfolding *
   27.16 -  proof (rule setsum_reindex_nonzero[symmetric])
   27.17 +  proof (rule setsum.reindex_nontrivial[symmetric])
   27.18      show "finite p"
   27.19        using assm by auto
   27.20      fix x y
   27.21 @@ -1932,7 +1932,7 @@
   27.22    assumes "content (cbox a b) = 0"
   27.23      and "p tagged_division_of (cbox a b)"
   27.24    shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
   27.25 -proof (rule setsum_0', rule)
   27.26 +proof (rule setsum.neutral, rule)
   27.27    fix y
   27.28    assume y: "y \<in> p"
   27.29    obtain x k where xk: "y = (x, k)"
   27.30 @@ -2541,7 +2541,7 @@
   27.31      proof -
   27.32        case goal1
   27.33        have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0"
   27.34 -      proof (rule setsum_0', rule)
   27.35 +      proof (rule setsum.neutral, rule)
   27.36          fix x
   27.37          assume x: "x \<in> p"
   27.38          have "f (fst x) = 0"
   27.39 @@ -2775,8 +2775,8 @@
   27.40          assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d1 x \<inter> d2 x) fine p"
   27.41          have *: "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) =
   27.42            (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)"
   27.43 -          unfolding scaleR_right_distrib setsum_addf[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,symmetric]
   27.44 -          by (rule setsum_cong2) auto
   27.45 +          unfolding scaleR_right_distrib setsum.distrib[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,symmetric]
   27.46 +          by (rule setsum.cong) auto
   27.47          have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) =
   27.48            norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
   27.49            unfolding * by (auto simp add: algebra_simps)
   27.50 @@ -2957,7 +2957,7 @@
   27.51  next
   27.52    case (insert x F)
   27.53    show ?case
   27.54 -    unfolding setsum_insert[OF insert(1,3)]
   27.55 +    unfolding setsum.insert[OF insert(1,3)]
   27.56      apply (rule has_integral_add)
   27.57      using insert assms
   27.58      apply auto
   27.59 @@ -3257,7 +3257,7 @@
   27.60      apply (subst *(1))
   27.61      defer
   27.62      apply (subst *(1))
   27.63 -    unfolding setprod_insert[OF *(2-)]
   27.64 +    unfolding setprod.insert[OF *(2-)]
   27.65      apply auto
   27.66      done
   27.67    assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
   27.68 @@ -3270,7 +3270,7 @@
   27.69        (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)"
   27.70      "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) =
   27.71        (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))"
   27.72 -    by (auto intro!: setprod_cong)
   27.73 +    by (auto intro!: setprod.cong)
   27.74    have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
   27.75      unfolding not_le
   27.76      using as[unfolded ,rule_format,of k] assms
   27.77 @@ -3539,7 +3539,7 @@
   27.78      have lem3: "\<And>g :: 'a set \<Rightarrow> 'a set. finite p \<Longrightarrow>
   27.79        setsum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} =
   27.80        setsum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)"
   27.81 -      apply (rule setsum_mono_zero_left)
   27.82 +      apply (rule setsum.mono_neutral_left)
   27.83        prefer 3
   27.84      proof
   27.85        fix g :: "'a set \<Rightarrow> 'a set"
   27.86 @@ -3664,9 +3664,9 @@
   27.87        also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
   27.88          (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
   27.89          unfolding lem3[OF p(3)]
   27.90 -        apply (subst setsum_reindex_nonzero[OF p(3)])
   27.91 +        apply (subst setsum.reindex_nontrivial[OF p(3)])
   27.92          defer
   27.93 -        apply (subst setsum_reindex_nonzero[OF p(3)])
   27.94 +        apply (subst setsum.reindex_nontrivial[OF p(3)])
   27.95          defer
   27.96          unfolding lem4[symmetric]
   27.97          apply (rule refl)
   27.98 @@ -3689,7 +3689,7 @@
   27.99            apply auto
  27.100            done
  27.101        qed
  27.102 -      also note setsum_addf[symmetric]
  27.103 +      also note setsum.distrib[symmetric]
  27.104        also have *: "\<And>x. x \<in> p \<Longrightarrow>
  27.105          (\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x +
  27.106            (\<lambda>(x,ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x =
  27.107 @@ -3705,7 +3705,7 @@
  27.108            unfolding uv content_split[OF k,of u v c]
  27.109            by auto
  27.110        qed
  27.111 -      note setsum_cong2[OF this]
  27.112 +      note setsum.cong [OF _ this]
  27.113        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 +
  27.114          ((\<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) =
  27.115          (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)"
  27.116 @@ -3774,7 +3774,7 @@
  27.117      note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
  27.118      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) =
  27.119        norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
  27.120 -      apply (subst setsum_Un_zero)
  27.121 +      apply (subst setsum.union_inter_neutral)
  27.122        apply (rule p1 p2)+
  27.123        apply rule
  27.124        unfolding split_paired_all split_conv
  27.125 @@ -3799,12 +3799,13 @@
  27.126            using e k by (auto simp: inner_simps inner_not_same_Basis)
  27.127          have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
  27.128            (\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))"
  27.129 -          apply (rule setsum_cong2)
  27.130 +          apply (rule setsum.cong)
  27.131 +          apply (rule refl)
  27.132            apply (subst *)
  27.133            apply auto
  27.134            done
  27.135          also have "\<dots> < e"
  27.136 -          apply (subst setsum_delta)
  27.137 +          apply (subst setsum.delta)
  27.138            using e
  27.139            apply auto
  27.140            done
  27.141 @@ -4841,7 +4842,7 @@
  27.142    shows "setsum f s = iterate op + s f"
  27.143  proof -
  27.144    have *: "setsum f s = setsum f (support op + f s)"
  27.145 -    apply (rule setsum_mono_zero_right)
  27.146 +    apply (rule setsum.mono_neutral_right)
  27.147      unfolding support_def neutral_add
  27.148      using assms
  27.149      apply auto
  27.150 @@ -4923,7 +4924,8 @@
  27.151    apply (rule mult_left_mono)
  27.152    apply (rule order_trans[of _ "setsum content p"])
  27.153    apply (rule eq_refl)
  27.154 -  apply (rule setsum_cong2)
  27.155 +  apply (rule setsum.cong)
  27.156 +  apply (rule refl)
  27.157    apply (subst abs_of_nonneg)
  27.158    unfolding additive_content_division[OF assms(1)]
  27.159  proof -
  27.160 @@ -4957,7 +4959,8 @@
  27.161      apply (rule mult_left_mono)
  27.162      apply (rule order_trans[of _ "setsum (content \<circ> snd) p"])
  27.163      apply (rule eq_refl)
  27.164 -    apply (rule setsum_cong2)
  27.165 +    apply (rule setsum.cong)
  27.166 +    apply (rule refl)
  27.167      apply (subst o_def)
  27.168      apply (rule abs_of_nonneg)
  27.169    proof -
  27.170 @@ -4995,7 +4998,7 @@
  27.171    apply (rule eq_refl)
  27.172    apply (rule arg_cong[where f=norm])
  27.173    unfolding setsum_subtractf[symmetric]
  27.174 -  apply (rule setsum_cong2)
  27.175 +  apply (rule setsum.cong)
  27.176    unfolding scaleR_diff_right
  27.177    apply auto
  27.178    done
  27.179 @@ -5516,7 +5519,7 @@
  27.180        (\<Prod>i\<in>Basis - {k}. interval_upperbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i -
  27.181          interval_lowerbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i) =
  27.182        (\<Prod>i\<in>Basis - {k}. b\<bullet>i - a\<bullet>i)"
  27.183 -      apply (rule setprod_cong)
  27.184 +      apply (rule setprod.cong)
  27.185        apply (rule refl)
  27.186        unfolding interval_doublesplit[OF k]
  27.187        apply (subst interval_bounds)
  27.188 @@ -5533,7 +5536,7 @@
  27.189        apply (rule assms)
  27.190        unfolding if_not_P
  27.191        apply (subst *)
  27.192 -      apply (subst setprod_insert)
  27.193 +      apply (subst setprod.insert)
  27.194        unfolding **
  27.195        unfolding interval_doublesplit[OF k] box_eq_empty not_ex not_less
  27.196        prefer 3
  27.197 @@ -5574,7 +5577,8 @@
  27.198      assume p: "p tagged_division_of (cbox a b) \<and> (\<lambda>x. ball x d) fine p"
  27.199      have *: "(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) =
  27.200        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. abs(x\<bullet>k - c) \<le> d}) *\<^sub>R ?i x)"
  27.201 -      apply (rule setsum_cong2)
  27.202 +      apply (rule setsum.cong)
  27.203 +      apply (rule refl)
  27.204        unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
  27.205        apply cases
  27.206        apply (rule disjI1)
  27.207 @@ -5857,14 +5861,14 @@
  27.208      (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
  27.209    show ?case
  27.210      unfolding *
  27.211 -    apply (subst setsum_Un_disjoint)
  27.212 -    unfolding setsum_insert[OF insert(1-2)]
  27.213 +    apply (subst setsum.union_disjoint)
  27.214 +    unfolding setsum.insert[OF insert(1-2)]
  27.215      prefer 4
  27.216      apply (subst insert(3))
  27.217      unfolding add_right_cancel
  27.218    proof -
  27.219      show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
  27.220 -      apply (subst setsum_reindex)
  27.221 +      apply (subst setsum.reindex)
  27.222        unfolding inj_on_def
  27.223        apply auto
  27.224        done
  27.225 @@ -6678,7 +6682,7 @@
  27.226      unfolding *
  27.227      apply (subst setsum_iterate[symmetric])
  27.228      defer
  27.229 -    apply (rule setsum_cong2)
  27.230 +    apply (rule setsum.cong)
  27.231      unfolding split_paired_all split_conv
  27.232      using assms(2)
  27.233      apply auto
  27.234 @@ -6933,7 +6937,7 @@
  27.235        have "norm (y - x) < e + setsum (\<lambda>i. 0) Basis"
  27.236          apply (rule le_less_trans[OF norm_le_l1])
  27.237          apply (subst *)
  27.238 -        apply (subst setsum_insert)
  27.239 +        apply (subst setsum.insert)
  27.240          prefer 3
  27.241          apply (rule add_less_le_mono)
  27.242        proof -
  27.243 @@ -7425,7 +7429,7 @@
  27.244            using assms(7)
  27.245            unfolding algebra_simps add_left_cancel scaleR_right.setsum
  27.246            by (subst setsum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
  27.247 -             (auto intro!: * setsum_cong simp: bij_betw_def dest!: p(4))
  27.248 +             (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4))
  27.249        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")
  27.250          unfolding scaleR_diff_right scaleR_scaleR
  27.251          using assms(1)
  27.252 @@ -7449,14 +7453,6 @@
  27.253    unfolding image_affinity_cbox
  27.254    by auto
  27.255  
  27.256 -lemma setprod_cong2:
  27.257 -  assumes "\<And>x. x \<in> A \<Longrightarrow> f x = g x"
  27.258 -  shows "setprod f A = setprod g A"
  27.259 -  apply (rule setprod_cong)
  27.260 -  using assms
  27.261 -  apply auto
  27.262 -  done
  27.263 -
  27.264  lemma content_image_affinity_cbox:
  27.265    "content((\<lambda>x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) =
  27.266      abs m ^ DIM('a) * content (cbox a b)" (is "?l = ?r")
  27.267 @@ -7486,7 +7482,7 @@
  27.268        by (simp add: inner_simps field_simps)
  27.269      ultimately show ?thesis
  27.270        by (simp add: image_affinity_cbox True content_cbox'
  27.271 -        setprod_timesf setprod_constant inner_diff_left)
  27.272 +        setprod.distrib setprod_constant inner_diff_left)
  27.273    next
  27.274      case False
  27.275      with as have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
  27.276 @@ -7499,7 +7495,7 @@
  27.277        by (simp add: inner_simps field_simps)
  27.278      ultimately show ?thesis using False
  27.279        by (simp add: image_affinity_cbox content_cbox'
  27.280 -        setprod_timesf[symmetric] setprod_constant[symmetric] inner_diff_left)
  27.281 +        setprod.distrib[symmetric] setprod_constant[symmetric] inner_diff_left)
  27.282    qed
  27.283  qed
  27.284  
  27.285 @@ -7593,8 +7589,9 @@
  27.286      unfolding content_def image_stretch_interval
  27.287      apply -
  27.288      unfolding interval_bounds' if_not_P
  27.289 -    unfolding abs_setprod setprod_timesf[symmetric]
  27.290 -    apply (rule setprod_cong2)
  27.291 +    unfolding abs_setprod setprod.distrib[symmetric]
  27.292 +    apply (rule setprod.cong)
  27.293 +    apply (rule refl)
  27.294      unfolding lessThan_iff
  27.295      apply (simp only: inner_setsum_left_Basis)
  27.296    proof -
  27.297 @@ -7928,7 +7925,7 @@
  27.298        unfolding setsum_right_distrib
  27.299        apply (subst(2) pA)
  27.300        apply (subst pA)
  27.301 -      unfolding setsum_Un_disjoint[OF pA(2-)]
  27.302 +      unfolding setsum.union_disjoint[OF pA(2-)]
  27.303      proof (rule norm_triangle_le, rule **)
  27.304        case goal1
  27.305        show ?case
  27.306 @@ -7993,7 +7990,7 @@
  27.307          apply rule
  27.308          apply (unfold split_paired_all split_conv)
  27.309          defer
  27.310 -        unfolding setsum_Un_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
  27.311 +        unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
  27.312          unfolding setsum_right_distrib[symmetric]
  27.313          thm additive_tagged_division_1
  27.314          apply (subst additive_tagged_division_1[OF _ as(1)])
  27.315 @@ -8013,7 +8010,7 @@
  27.316          show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x -
  27.317            (f ((Sup k)) - f ((Inf k)))) \<le> e * (b - a) / 2"
  27.318            apply (rule *[where t="p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0}"])
  27.319 -          apply (rule setsum_mono_zero_right[OF pA(2)])
  27.320 +          apply (rule setsum.mono_neutral_right[OF pA(2)])
  27.321            defer
  27.322            apply rule
  27.323            unfolding split_paired_all split_conv o_def
  27.324 @@ -8049,7 +8046,7 @@
  27.325            case goal2
  27.326            show ?case
  27.327              apply (subst *)
  27.328 -            apply (subst setsum_Un_disjoint)
  27.329 +            apply (subst setsum.union_disjoint)
  27.330              prefer 4
  27.331              apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"])
  27.332              apply (rule norm_triangle_le,rule add_mono)
  27.333 @@ -8480,7 +8477,7 @@
  27.334        then show ?thesis
  27.335          unfolding ** box_real
  27.336          apply -
  27.337 -        apply (subst setsum_insert)
  27.338 +        apply (subst setsum.insert)
  27.339          apply (rule p')
  27.340          unfolding split_conv
  27.341          defer
  27.342 @@ -9855,11 +9852,12 @@
  27.343          unfolding if_P[OF True]
  27.344          apply (rule trans)
  27.345          defer
  27.346 -        apply (rule setsum_cong2)
  27.347 +        apply (rule setsum.cong)
  27.348 +        apply (rule refl)
  27.349          apply (subst *)
  27.350          apply assumption
  27.351          apply (rule refl)
  27.352 -        unfolding setsum_delta[OF assms(1)]
  27.353 +        unfolding setsum.delta[OF assms(1)]
  27.354          using s
  27.355          apply auto
  27.356          done
  27.357 @@ -9995,7 +9993,7 @@
  27.358      apply (rule setsum_over_tagged_division_lemma[OF assms(1)])
  27.359      apply (rule integral_null)
  27.360      apply assumption
  27.361 -    apply (rule setsum_cong2)
  27.362 +    apply (rule setsum.cong)
  27.363      using assms(2)
  27.364      apply auto
  27.365      done
  27.366 @@ -10153,7 +10151,7 @@
  27.367  
  27.368    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) -
  27.369      integral (cbox a b) f) < e"
  27.370 -    apply (subst setsum_Un_zero[symmetric])
  27.371 +    apply (subst setsum.union_inter_neutral[symmetric])
  27.372      apply (rule p')
  27.373      prefer 3
  27.374      apply assumption
  27.375 @@ -10180,13 +10178,9 @@
  27.376  
  27.377    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))
  27.378      (qq ` r) - integral (cbox a b) f) < e"
  27.379 -    apply (subst (asm) setsum_UNION_zero)
  27.380 -    prefer 4
  27.381 -    apply assumption
  27.382 -    apply (rule finite_imageI)
  27.383 -    apply fact
  27.384 +    apply (subst (asm) setsum.Union_comp)
  27.385 +    prefer 2
  27.386      unfolding split_paired_all split_conv image_iff
  27.387 -    defer
  27.388      apply (erule bexE)+
  27.389    proof -
  27.390      fix x m k l T1 T2
  27.391 @@ -10215,9 +10209,9 @@
  27.392  
  27.393    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 -
  27.394      integral (cbox a b) f) < e"
  27.395 -    apply (subst (asm) setsum_reindex_nonzero)
  27.396 +    apply (subst (asm) setsum.reindex_nontrivial)
  27.397      apply fact
  27.398 -    apply (rule setsum_0')
  27.399 +    apply (rule setsum.neutral)
  27.400      apply rule
  27.401      unfolding split_paired_all split_conv
  27.402      defer
  27.403 @@ -10247,7 +10241,7 @@
  27.404    proof -
  27.405      case goal2
  27.406      have *: "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)"
  27.407 -      apply (subst setsum_reindex_nonzero)
  27.408 +      apply (subst setsum.reindex_nontrivial)
  27.409        apply fact
  27.410        unfolding split_paired_all snd_conv split_def o_def
  27.411      proof -
  27.412 @@ -10264,12 +10258,11 @@
  27.413          apply auto
  27.414          done
  27.415      qed auto
  27.416 +    from q(1) have **: "snd ` p \<union> q = q" by auto
  27.417      show ?case
  27.418        unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def
  27.419 -      apply (rule setsum_Un_disjoint'[symmetric])
  27.420 -      using q(1) q'(1) p'(1)
  27.421 -      apply auto
  27.422 -      done
  27.423 +      using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\<lambda>k. integral k f", symmetric]
  27.424 +        by simp
  27.425    next
  27.426      case goal1
  27.427      have *: "k * real (card r) / (1 + real (card r)) < k"
  27.428 @@ -10377,7 +10370,7 @@
  27.429      done
  27.430    have th: "op ^ x o op + m = (\<lambda>i. x^m * x^i)"
  27.431      by (rule ext) (simp add: power_add power_mult)
  27.432 -  from setsum_reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
  27.433 +  from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
  27.434    have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})"
  27.435      by simp
  27.436    then show ?thesis
  27.437 @@ -10622,7 +10615,8 @@
  27.438                  [of _ "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \<in> p. m (fst xk) = t})"])
  27.439                apply (rule eq_refl)
  27.440                apply (rule arg_cong[where f=norm])
  27.441 -              apply (rule setsum_cong2)
  27.442 +              apply (rule setsum.cong)
  27.443 +              apply (rule refl)
  27.444                defer
  27.445                apply (rule henstock_lemma_part1)
  27.446                apply (rule assms(1)[rule_format])
  27.447 @@ -11507,7 +11501,7 @@
  27.448              apply auto
  27.449              done
  27.450            also have "\<dots> = (\<Sum>k\<in>{k \<inter> l |l. l \<in> snd ` p}. norm (integral k f))"
  27.451 -            apply (rule setsum_mono_zero_left)
  27.452 +            apply (rule setsum.mono_neutral_left)
  27.453              apply (subst simple_image)
  27.454              apply (rule finite_imageI)+
  27.455              apply fact
  27.456 @@ -11525,7 +11519,7 @@
  27.457            qed
  27.458            also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (k \<inter> l) f))"
  27.459              unfolding simple_image
  27.460 -            apply (rule setsum_reindex_nonzero[unfolded o_def])
  27.461 +            apply (rule setsum.reindex_nontrivial [unfolded o_def])
  27.462              apply (rule finite_imageI)
  27.463              apply (rule p')
  27.464            proof -
  27.465 @@ -11557,7 +11551,7 @@
  27.466            unfolding split_def ..
  27.467          also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
  27.468            unfolding *
  27.469 -          apply (rule setsum_reindex_nonzero[symmetric,unfolded o_def])
  27.470 +          apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def])
  27.471            apply (rule finite_product_dependent)
  27.472            apply fact
  27.473            apply (rule finite_imageI)
  27.474 @@ -11598,7 +11592,7 @@
  27.475          qed
  27.476          also have "\<dots> = (\<Sum>(x, k)\<in>p'. norm (integral k f))"
  27.477            unfolding sum_p'
  27.478 -          apply (rule setsum_mono_zero_right)
  27.479 +          apply (rule setsum.mono_neutral_right)
  27.480            apply (subst *)
  27.481            apply (rule finite_imageI[OF finite_product_dependent])
  27.482            apply fact
  27.483 @@ -11642,7 +11636,7 @@
  27.484          note pdfin = finite_cartesian_product[OF p'(1) d'(1)]
  27.485          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))"
  27.486            unfolding norm_scaleR
  27.487 -          apply (rule setsum_mono_zero_left)
  27.488 +          apply (rule setsum.mono_neutral_left)
  27.489            apply (subst *)
  27.490            apply (rule finite_imageI)
  27.491            apply fact
  27.492 @@ -11656,7 +11650,7 @@
  27.493            done
  27.494          also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
  27.495            unfolding *
  27.496 -          apply (subst setsum_reindex_nonzero)
  27.497 +          apply (subst setsum.reindex_nontrivial)
  27.498            apply fact
  27.499            unfolding split_paired_all
  27.500            unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff Pair_eq
  27.501 @@ -11694,7 +11688,8 @@
  27.502            apply (rule p')
  27.503            apply rule
  27.504            apply (rule d')
  27.505 -          apply (rule setsum_cong2)
  27.506 +          apply (rule setsum.cong)
  27.507 +          apply (rule refl)
  27.508            unfolding split_paired_all split_conv
  27.509          proof -
  27.510            fix x l
  27.511 @@ -11702,7 +11697,8 @@
  27.512            note xl = p'(2-4)[OF this]
  27.513            from this(3) guess u v by (elim exE) note uv=this
  27.514            have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
  27.515 -            apply (rule setsum_cong2)
  27.516 +            apply (rule setsum.cong)
  27.517 +            apply (rule refl)
  27.518              apply (drule d'(4))
  27.519              apply safe
  27.520              apply (subst Int_commute)
  27.521 @@ -11712,7 +11708,7 @@
  27.522              done
  27.523            also have "\<dots> = setsum content {k \<inter> cbox u v| k. k \<in> d}"
  27.524              unfolding simple_image
  27.525 -            apply (rule setsum_reindex_nonzero[unfolded o_def,symmetric])
  27.526 +            apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric])
  27.527              apply (rule d')
  27.528            proof -
  27.529              case goal1
  27.530 @@ -11731,7 +11727,7 @@
  27.531                unfolding uv inter_interval content_eq_0_interior ..
  27.532            qed
  27.533            also have "\<dots> = setsum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
  27.534 -            apply (rule setsum_mono_zero_right)
  27.535 +            apply (rule setsum.mono_neutral_right)
  27.536              unfolding simple_image
  27.537              apply (rule finite_imageI)
  27.538              apply (rule d')
  27.539 @@ -11901,7 +11897,8 @@
  27.540              using d1(2)[rule_format,OF conjI[OF p(1,2)]]
  27.541              by (simp only: real_norm_def)
  27.542            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))"
  27.543 -            apply (rule setsum_cong2)
  27.544 +            apply (rule setsum.cong)
  27.545 +            apply (rule refl)
  27.546              unfolding split_paired_all split_conv
  27.547              apply (drule tagged_division_ofD(4)[OF p(1)])
  27.548              unfolding norm_scaleR
  27.549 @@ -11965,7 +11962,7 @@
  27.550      then have "(\<Sum>k\<in>d. norm (integral k (\<lambda>x. f x + g x))) \<le>
  27.551        (\<Sum>k\<in>d. norm (integral k f)) + (\<Sum>k\<in>d. norm (integral k g))"
  27.552        apply -
  27.553 -      unfolding setsum_addf[symmetric]
  27.554 +      unfolding setsum.distrib [symmetric]
  27.555        apply (rule setsum_mono)
  27.556        apply (subst integral_add)
  27.557        prefer 3
  27.558 @@ -12080,7 +12077,7 @@
  27.559    have *: "\<And>x. ?Tf x = (\<Sum>i\<in>Basis.
  27.560      ((\<lambda>y. (\<Sum>j\<in>Basis. (if j = i then y else 0) *\<^sub>R j)) o
  27.561       (\<lambda>x. (norm (\<Sum>j\<in>Basis. (if j = i then f x\<bullet>T i else 0) *\<^sub>R j)))) x)"
  27.562 -    by (simp add: comp_def if_distrib setsum_cases)
  27.563 +    by (simp add: comp_def if_distrib setsum.If_cases)
  27.564    show ?thesis
  27.565      unfolding *
  27.566      apply (rule absolutely_integrable_setsum[OF finite_Basis])
  27.567 @@ -12191,7 +12188,7 @@
  27.568          done
  27.569      qed
  27.570      also have "\<dots> \<le> setsum (op \<bullet> (integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis"
  27.571 -      apply (subst setsum_commute)
  27.572 +      apply (subst setsum.commute)
  27.573        apply (rule setsum_mono)
  27.574      proof -
  27.575        case goal1
    28.1 --- a/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Fri Jun 27 22:08:55 2014 +0200
    28.2 +++ b/src/HOL/Multivariate_Analysis/Linear_Algebra.thy	Sat Jun 28 09:16:42 2014 +0200
    28.3 @@ -166,8 +166,8 @@
    28.4    assumes fS: "finite S" and fT: "finite T" and fST: "f ` S \<subseteq> T"
    28.5    shows "setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) T = setsum g S"
    28.6    apply (subst setsum_image_gen[OF fS, of g f])
    28.7 -  apply (rule setsum_mono_zero_right[OF fT fST])
    28.8 -  apply (auto intro: setsum_0')
    28.9 +  apply (rule setsum.mono_neutral_right[OF fT fST])
   28.10 +  apply (auto intro: setsum.neutral)
   28.11    done
   28.12  
   28.13  lemma vector_eq_ldot: "(\<forall>x. x \<bullet> y = x \<bullet> z) \<longleftrightarrow> y = z"
   28.14 @@ -370,13 +370,13 @@
   28.15      apply (auto simp add: bilinear_def)
   28.16      done
   28.17    also have "\<dots> = setsum (\<lambda>x. setsum (\<lambda>y. h (f x) (g y)) T) S"
   28.18 -    apply (rule setsum_cong, simp)
   28.19 +    apply (rule setsum.cong, simp)
   28.20      apply (rule linear_setsum[unfolded o_def])
   28.21      using bh fT
   28.22      apply (auto simp add: bilinear_def)
   28.23      done
   28.24    finally show ?thesis
   28.25 -    unfolding setsum_cartesian_product .
   28.26 +    unfolding setsum.cartesian_product .
   28.27  qed
   28.28  
   28.29  
   28.30 @@ -1090,7 +1090,7 @@
   28.31        assume xS: "x \<notin> S"
   28.32        have th00: "(\<Sum>v\<in>S. (if v = x then c else u v) *\<^sub>R v) = y"
   28.33          unfolding u[symmetric]
   28.34 -        apply (rule setsum_cong2)
   28.35 +        apply (rule setsum.cong)
   28.36          using xS
   28.37          apply auto
   28.38          done
   28.39 @@ -1123,7 +1123,7 @@
   28.40        using fS aS
   28.41        apply simp
   28.42        apply (subst (2) ua[symmetric])
   28.43 -      apply (rule setsum_cong2)
   28.44 +      apply (rule setsum.cong)
   28.45        apply auto
   28.46        done
   28.47      with th0 have ?rhs by fast
   28.48 @@ -1174,7 +1174,7 @@
   28.49        unfolding span_explicit by blast
   28.50      let ?u = "\<lambda>x. if x \<in> S' then u x else 0"
   28.51      have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = setsum (\<lambda>v. u v *\<^sub>R v) S'"
   28.52 -      using SS' fS by (auto intro!: setsum_mono_zero_cong_right)
   28.53 +      using SS' fS by (auto intro!: setsum.mono_neutral_cong_right)
   28.54      then have "setsum (\<lambda>v. ?u v *\<^sub>R v) S = y" by (metis u)
   28.55      then have "y \<in> ?rhs" by auto
   28.56    }
   28.57 @@ -1428,14 +1428,14 @@
   28.58    have "(\<Sum>x\<in>P. norm (f x)) \<le> (\<Sum>x\<in>P. \<Sum>b\<in>Basis. \<bar>f x \<bullet> b\<bar>)"
   28.59      by (rule setsum_mono) (rule norm_le_l1)
   28.60    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>)"
   28.61 -    by (rule setsum_commute)
   28.62 +    by (rule setsum.commute)
   28.63    also have "\<dots> \<le> of_nat (card (Basis :: 'n set)) * (2 * e)"
   28.64    proof (rule setsum_bounded)
   28.65      fix i :: 'n
   28.66      assume i: "i \<in> Basis"
   28.67      have "norm (\<Sum>x\<in>P. \<bar>f x \<bullet> i\<bar>) \<le>
   28.68        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)"
   28.69 -      by (simp add: abs_real_def setsum_cases[OF fP] setsum_negf norm_triangle_ineq4 inner_setsum_left
   28.70 +      by (simp add: abs_real_def setsum.If_cases[OF fP] setsum_negf norm_triangle_ineq4 inner_setsum_left
   28.71          del: real_norm_def)
   28.72      also have "\<dots> \<le> e + e"
   28.73        unfolding real_norm_def
   28.74 @@ -1538,7 +1538,7 @@
   28.75      unfolding bilinear_setsum[OF bh finite_Basis finite_Basis] ..
   28.76    finally have th: "norm (h x y) = \<dots>" .
   28.77    show "norm (h x y) \<le> (\<Sum>i\<in>Basis. \<Sum>j\<in>Basis. norm (h i j)) * norm x * norm y"
   28.78 -    apply (auto simp add: setsum_left_distrib th setsum_cartesian_product)
   28.79 +    apply (auto simp add: setsum_left_distrib th setsum.cartesian_product)
   28.80      apply (rule setsum_norm_le)
   28.81      apply simp
   28.82      apply (auto simp add: bilinear_rmul[OF bh] bilinear_lmul[OF bh]
   28.83 @@ -1958,9 +1958,9 @@
   28.84      have "orthogonal ?a y"
   28.85        unfolding orthogonal_def
   28.86        unfolding inner_diff inner_setsum_left right_minus_eq
   28.87 -      unfolding setsum_diff1' [OF `finite C` `y \<in> C`]
   28.88 +      unfolding setsum.remove [OF `finite C` `y \<in> C`]
   28.89        apply (clarsimp simp add: inner_commute[of y a])
   28.90 -      apply (rule setsum_0')
   28.91 +      apply (rule setsum.neutral)
   28.92        apply clarsimp
   28.93        apply (rule C(4)[unfolded pairwise_def orthogonal_def, rule_format])
   28.94        using `y \<in> C` by auto
   28.95 @@ -2052,7 +2052,7 @@
   28.96          unfolding setsum_clauses(2)[OF fth]
   28.97          apply simp unfolding inner_simps
   28.98          apply (clarsimp simp add: inner_add inner_setsum_left)
   28.99 -        apply (rule setsum_0', rule ballI)
  28.100 +        apply (rule setsum.neutral, rule ballI)
  28.101          unfolding inner_commute
  28.102          apply (auto simp add: x field_simps
  28.103            intro: B(5)[unfolded pairwise_def orthogonal_def, rule_format])
  28.104 @@ -2806,8 +2806,8 @@
  28.105    by (simp add: Basis_le_norm infnorm_Max)
  28.106  
  28.107  lemma (in euclidean_space) euclidean_inner: "inner x y = (\<Sum>b\<in>Basis. (x \<bullet> b) * (y \<bullet> b))"
  28.108 -  by (subst (1 2) euclidean_representation[symmetric])
  28.109 -    (simp add: inner_setsum_left inner_setsum_right setsum_cases inner_Basis ac_simps if_distrib)
  28.110 +  by (subst (1 2) euclidean_representation [symmetric])
  28.111 +    (simp add: inner_setsum_right inner_Basis ac_simps)
  28.112  
  28.113  lemma norm_le_infnorm:
  28.114    fixes x :: "'a::euclidean_space"
    29.1 --- a/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    29.2 +++ b/src/HOL/Multivariate_Analysis/Ordered_Euclidean_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    29.3 @@ -48,7 +48,7 @@
    29.4  
    29.5  lemma inner_Basis_inf_left: "i \<in> Basis \<Longrightarrow> inf x y \<bullet> i = inf (x \<bullet> i) (y \<bullet> i)"
    29.6    and inner_Basis_sup_left: "i \<in> Basis \<Longrightarrow> sup x y \<bullet> i = sup (x \<bullet> i) (y \<bullet> i)"
    29.7 -  by (simp_all add: eucl_inf eucl_sup inner_setsum_left inner_Basis if_distrib setsum_delta
    29.8 +  by (simp_all add: eucl_inf eucl_sup inner_setsum_left inner_Basis if_distrib comm_monoid_add_class.setsum.delta
    29.9        cong: if_cong)
   29.10  
   29.11  lemma inner_Basis_INF_left: "i \<in> Basis \<Longrightarrow> (INF x:X. f x) \<bullet> i = (INF x:X. f x \<bullet> i)"
   29.12 @@ -119,7 +119,7 @@
   29.13    hence "Inf ?proj = x \<bullet> b"
   29.14      by (auto intro!: conditionally_complete_lattice_class.cInf_eq_minimum simp del: Inf_class.Inf_image_eq)
   29.15    hence "x \<bullet> b = Inf X \<bullet> b"
   29.16 -    by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum_delta
   29.17 +    by (auto simp: eucl_Inf Inf_class.INF_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum.delta
   29.18        simp del: Inf_class.Inf_image_eq
   29.19        cong: if_cong)
   29.20    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
   29.21 @@ -142,7 +142,7 @@
   29.22    hence "Sup ?proj = x \<bullet> b"
   29.23      by (auto intro!: cSup_eq_maximum simp del: Sup_image_eq)
   29.24    hence "x \<bullet> b = Sup X \<bullet> b"
   29.25 -    by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum_delta
   29.26 +    by (auto simp: eucl_Sup[where 'a='a] SUP_def inner_setsum_left inner_Basis if_distrib `b \<in> Basis` setsum.delta
   29.27        simp del: Sup_image_eq cong: if_cong)
   29.28    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
   29.29  qed
    30.1 --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    30.2 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    30.3 @@ -7173,7 +7173,7 @@
    30.4      then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d"
    30.5        by (simp add: span_setsum span_clauses)
    30.6      also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)"
    30.7 -      by (rule setsum_mono_zero_cong_left [OF finite_Basis d]) (auto simp add: x)
    30.8 +      by (rule setsum.mono_neutral_cong_left [OF finite_Basis d]) (auto simp add: x)
    30.9      finally show "x \<in> span d"
   30.10        unfolding euclidean_representation .
   30.11    qed
    31.1 --- a/src/HOL/NSA/HSeries.thy	Fri Jun 27 22:08:55 2014 +0200
    31.2 +++ b/src/HOL/NSA/HSeries.thy	Sat Jun 28 09:16:42 2014 +0200
    31.3 @@ -55,7 +55,7 @@
    31.4  
    31.5  lemma sumhr_add:
    31.6    "!!m n. sumhr (m,n,f) + sumhr(m,n,g) = sumhr(m,n,%i. f i + g i)"
    31.7 -unfolding sumhr_app by transfer (rule setsum_addf [symmetric])
    31.8 +unfolding sumhr_app by transfer (rule setsum.distrib [symmetric])
    31.9  
   31.10  lemma sumhr_mult:
   31.11    "!!m n. hypreal_of_real r * sumhr(m,n,f) = sumhr(m,n,%n. r * f n)"
    32.1 --- a/src/HOL/Nat_Transfer.thy	Fri Jun 27 22:08:55 2014 +0200
    32.2 +++ b/src/HOL/Nat_Transfer.thy	Sat Jun 28 09:16:42 2014 +0200
    32.3 @@ -194,9 +194,9 @@
    32.4  lemma transfer_nat_int_sum_prod:
    32.5      "setsum f A = setsum (%x. f (nat x)) (int ` A)"
    32.6      "setprod f A = setprod (%x. f (nat x)) (int ` A)"
    32.7 -  apply (subst setsum_reindex)
    32.8 +  apply (subst setsum.reindex)
    32.9    apply (unfold inj_on_def, auto)
   32.10 -  apply (subst setprod_reindex)
   32.11 +  apply (subst setprod.reindex)
   32.12    apply (unfold inj_on_def o_def, auto)
   32.13  done
   32.14  
   32.15 @@ -237,7 +237,7 @@
   32.16  *)
   32.17  
   32.18  (* Making A = B in this lemma doesn't work. Why not?
   32.19 -   Also, why aren't setsum_cong and setprod_cong enough,
   32.20 +   Also, why aren't setsum.cong and setprod.cong enough,
   32.21     with the previously mentioned rule turned on? *)
   32.22  
   32.23  lemma transfer_nat_int_sum_prod_cong:
   32.24 @@ -246,9 +246,9 @@
   32.25      "A = B \<Longrightarrow> nat_set B \<Longrightarrow> (!!x. x >= 0 \<Longrightarrow> f x = g x) \<Longrightarrow>
   32.26        setprod f A = setprod g B"
   32.27    unfolding nat_set_def
   32.28 -  apply (subst setsum_cong, assumption)
   32.29 +  apply (subst setsum.cong, assumption)
   32.30    apply auto [2]
   32.31 -  apply (subst setprod_cong, assumption, auto)
   32.32 +  apply (subst setprod.cong, assumption, auto)
   32.33  done
   32.34  
   32.35  declare transfer_morphism_nat_int [transfer add
   32.36 @@ -399,11 +399,11 @@
   32.37  lemma transfer_int_nat_sum_prod:
   32.38      "nat_set A \<Longrightarrow> setsum f A = setsum (%x. f (int x)) (nat ` A)"
   32.39      "nat_set A \<Longrightarrow> setprod f A = setprod (%x. f (int x)) (nat ` A)"
   32.40 -  apply (subst setsum_reindex)
   32.41 +  apply (subst setsum.reindex)
   32.42    apply (unfold inj_on_def nat_set_def, auto simp add: eq_nat_nat_iff)
   32.43 -  apply (subst setprod_reindex)
   32.44 +  apply (subst setprod.reindex)
   32.45    apply (unfold inj_on_def nat_set_def o_def, auto simp add: eq_nat_nat_iff
   32.46 -            cong: setprod_cong)
   32.47 +            cong: setprod.cong)
   32.48  done
   32.49  
   32.50  (* this handles the case where the *range* of f is int *)
   32.51 @@ -413,11 +413,11 @@
   32.52        setprod f A = int(setprod (%x. nat (f x)) A)"
   32.53    unfolding is_nat_def
   32.54    apply (subst int_setsum, auto)
   32.55 -  apply (subst int_setprod, auto simp add: cong: setprod_cong)
   32.56 +  apply (subst int_setprod, auto simp add: cong: setprod.cong)
   32.57  done
   32.58  
   32.59  declare transfer_morphism_int_nat [transfer add
   32.60    return: transfer_int_nat_sum_prod transfer_int_nat_sum_prod2
   32.61 -  cong: setsum_cong setprod_cong]
   32.62 +  cong: setsum.cong setprod.cong]
   32.63  
   32.64  end
    33.1 --- a/src/HOL/Number_Theory/Binomial.thy	Fri Jun 27 22:08:55 2014 +0200
    33.2 +++ b/src/HOL/Number_Theory/Binomial.thy	Sat Jun 28 09:16:42 2014 +0200
    33.3 @@ -173,7 +173,7 @@
    33.4    also have
    33.5        "\<dots> = a^(n+1) + b^(n+1) + 
    33.6              (\<Sum>k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))"
    33.7 -    by (auto simp add: field_simps setsum_addf [symmetric] choose_reduce_nat)
    33.8 +    by (auto simp add: field_simps setsum.distrib [symmetric] choose_reduce_nat)
    33.9    also have "\<dots> = (\<Sum>k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))"
   33.10      using decomp by (simp add: field_simps)
   33.11    finally show "?P (Suc n)" by simp
   33.12 @@ -263,13 +263,13 @@
   33.13    have *: "finite {1 .. n}" "0 \<notin> {1 .. n}" by auto
   33.14    have eq: "insert 0 {1 .. n} = {0..n}" by auto
   33.15    have **: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) = (\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
   33.16 -    apply (rule setprod_reindex_cong [where f = Suc])
   33.17 +    apply (rule setprod.reindex_cong [where l = Suc])
   33.18      using False
   33.19      apply (auto simp add: fun_eq_iff field_simps)
   33.20      done
   33.21    show ?thesis
   33.22      apply (simp add: pochhammer_def)
   33.23 -    unfolding setprod_insert [OF *, unfolded eq]
   33.24 +    unfolding setprod.insert [OF *, unfolded eq]
   33.25      using ** apply (simp add: field_simps)
   33.26      done
   33.27  qed
   33.28 @@ -278,7 +278,7 @@
   33.29    unfolding fact_altdef_nat
   33.30    apply (cases n)
   33.31     apply (simp_all add: of_nat_setprod pochhammer_Suc_setprod)
   33.32 -  apply (rule setprod_reindex_cong[where f=Suc])
   33.33 +  apply (rule setprod.reindex_cong [where l = Suc])
   33.34      apply (auto simp add: fun_eq_iff)
   33.35    done
   33.36  
   33.37 @@ -347,7 +347,7 @@
   33.38      using setprod_constant[where A="{0 .. h}" and y="- 1 :: 'a"]
   33.39      by auto
   33.40    show ?thesis
   33.41 -    unfolding Suc pochhammer_Suc_setprod eq setprod_timesf[symmetric]
   33.42 +    unfolding Suc pochhammer_Suc_setprod eq setprod.distrib[symmetric]
   33.43      by (rule setprod.reindex_bij_witness[where i="op - h" and j="op - h"])
   33.44         (auto simp: of_nat_diff)
   33.45  qed
   33.46 @@ -390,7 +390,7 @@
   33.47      by auto
   33.48    from False show ?thesis
   33.49      by (simp add: pochhammer_def gbinomial_def field_simps
   33.50 -      eq setprod_timesf[symmetric])
   33.51 +      eq setprod.distrib[symmetric])
   33.52  qed
   33.53  
   33.54  lemma binomial_fact_lemma: "k \<le> n \<Longrightarrow> fact k * fact (n - k) * (n choose k) = fact n"
   33.55 @@ -459,10 +459,10 @@
   33.56        apply (simp add: binomial_fact[OF kn, where ?'a = 'a]
   33.57          gbinomial_pochhammer field_simps pochhammer_Suc_setprod)
   33.58        apply (simp add: pochhammer_Suc_setprod fact_altdef_nat h
   33.59 -        of_nat_setprod setprod_timesf[symmetric] eq' del: One_nat_def power_Suc)
   33.60 -      unfolding setprod_Un_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
   33.61 +        of_nat_setprod setprod.distrib[symmetric] eq' del: One_nat_def power_Suc)
   33.62 +      unfolding setprod.union_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h]
   33.63        unfolding mult_assoc[symmetric]
   33.64 -      unfolding setprod_timesf[symmetric]
   33.65 +      unfolding setprod.distrib[symmetric]
   33.66        apply simp
   33.67        apply (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"])
   33.68        apply (auto simp: of_nat_diff)
   33.69 @@ -520,7 +520,7 @@
   33.70  next
   33.71    case (Suc h)
   33.72    have eq0: "(\<Prod>i\<in>{1..k}. (a + 1) - of_nat i) = (\<Prod>i\<in>{0..h}. a - of_nat i)"
   33.73 -    apply (rule strong_setprod_reindex_cong[where f = Suc])
   33.74 +    apply (rule setprod.reindex_cong [where l = Suc])
   33.75        using Suc
   33.76        apply auto
   33.77      done
   33.78 @@ -573,7 +573,7 @@
   33.79      by (auto simp: gr0_conv_Suc lessThan_Suc_atMost atLeast0AtMost)
   33.80    also have "\<dots> = (\<Prod>i<k. of_nat (n - i) / of_nat (k - i) :: 'a)"
   33.81      using `k \<le> n` unfolding fact_eq_rev_setprod_nat of_nat_setprod
   33.82 -    by (auto simp add: setprod_dividef intro!: setprod_cong of_nat_diff[symmetric])
   33.83 +    by (auto simp add: setprod_dividef intro!: setprod.cong of_nat_diff[symmetric])
   33.84    finally show ?thesis .
   33.85  next
   33.86    case False
   33.87 @@ -702,15 +702,15 @@
   33.88    also have "\<dots> = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (\<Sum>_\<in>\<Inter>I. -1 ^ (card I + 1)))" (is "_ = nat ?rhs")
   33.89      by(subst setsum_right_distrib) simp
   33.90    also have "?rhs = (\<Sum>(I, _)\<in>Sigma {I. I \<subseteq> A \<and> I \<noteq> {}} Inter. -1 ^ (card I + 1))"
   33.91 -    using assms by(subst setsum_Sigma)(auto)
   33.92 +    using assms by(subst setsum.Sigma)(auto)
   33.93    also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:UNIV. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). -1 ^ (card I + 1))"
   33.94 -    by(rule setsum_reindex_cong[where f="\<lambda>(x, y). (y, x)"])(auto intro: inj_onI simp add: split_beta)
   33.95 +    by (rule setsum.reindex_cong [where l = "\<lambda>(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta)
   33.96    also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:\<Union>A. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). -1 ^ (card I + 1))"
   33.97 -    using assms by(auto intro!: setsum_mono_zero_cong_right finite_SigmaI2 intro: finite_subset[where B="\<Union>A"])
   33.98 +    using assms by(auto intro!: setsum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\<Union>A"])
   33.99    also have "\<dots> = (\<Sum>x\<in>\<Union>A. (\<Sum>I|I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I. -1 ^ (card I + 1)))" 
  33.100 -    using assms by(subst setsum_Sigma) auto
  33.101 +    using assms by(subst setsum.Sigma) auto
  33.102    also have "\<dots> = (\<Sum>_\<in>\<Union>A. 1)" (is "setsum ?lhs _ = _")
  33.103 -  proof(rule setsum_cong[OF refl])
  33.104 +  proof(rule setsum.cong[OF refl])
  33.105      fix x
  33.106      assume x: "x \<in> \<Union>A"
  33.107      def K \<equiv> "{X \<in> A. x \<in> X}"
  33.108 @@ -723,13 +723,13 @@
  33.109          simp add: card_gt_0_iff[folded Suc_le_eq]
  33.110          dest: finite_subset intro: card_mono)
  33.111      ultimately have "?lhs x = (\<Sum>(i, I)\<in>(SIGMA i:{1..card A}. ?I i). -1 ^ (i + 1))"
  33.112 -      by(rule setsum_reindex_cong[where f=snd]) fastforce
  33.113 +      by (rule setsum.reindex_cong [where l = snd]) fastforce
  33.114      also have "\<dots> = (\<Sum>i=1..card A. (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. -1 ^ (i + 1)))"
  33.115 -      using assms by(subst setsum_Sigma) auto
  33.116 +      using assms by(subst setsum.Sigma) auto
  33.117      also have "\<dots> = (\<Sum>i=1..card A. -1 ^ (i + 1) * (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1))"
  33.118        by(subst setsum_right_distrib) simp
  33.119      also have "\<dots> = (\<Sum>i=1..card K. -1 ^ (i + 1) * (\<Sum>I|I \<subseteq> K \<and> card I = i. 1))" (is "_ = ?rhs")
  33.120 -    proof(rule setsum_mono_zero_cong_right[rule_format])
  33.121 +    proof(rule setsum.mono_neutral_cong_right[rule_format])
  33.122        show "{1..card K} \<subseteq> {1..card A}" using `finite A`
  33.123          by(auto simp add: K_def intro: card_mono)
  33.124      next
  33.125 @@ -746,7 +746,7 @@
  33.126        fix i
  33.127        have "(\<Sum>I | I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1) = (\<Sum>I | I \<subseteq> K \<and> card I = i. 1 :: int)"
  33.128          (is "?lhs = ?rhs")
  33.129 -        by(rule setsum_cong)(auto simp add: K_def)
  33.130 +        by(rule setsum.cong)(auto simp add: K_def)
  33.131        thus "-1 ^ (i + 1) * ?lhs = -1 ^ (i + 1) * ?rhs" by simp
  33.132      qed simp
  33.133      also have "{I. I \<subseteq> K \<and> card I = 0} = {{}}" using assms
    34.1 --- a/src/HOL/Number_Theory/Cong.thy	Fri Jun 27 22:08:55 2014 +0200
    34.2 +++ b/src/HOL/Number_Theory/Cong.thy	Sat Jun 28 09:16:42 2014 +0200
    34.3 @@ -805,7 +805,7 @@
    34.4      proof -
    34.5        from fin a have "?x = (SUM j:{i}. u j * b j) +
    34.6            (SUM j:A-{i}. u j * b j)"
    34.7 -        by (subst setsum_Un_disjoint [symmetric], auto intro: setsum_cong)
    34.8 +        by (subst setsum.union_disjoint [symmetric], auto intro: setsum.cong)
    34.9        then have "[?x = u i * b i + (SUM j:A-{i}. u j * b j)] (mod m i)"
   34.10          by auto
   34.11        also have "[u i * b i + (SUM j:A-{i}. u j * b j) =
    35.1 --- a/src/HOL/Number_Theory/Gauss.thy	Fri Jun 27 22:08:55 2014 +0200
    35.2 +++ b/src/HOL/Number_Theory/Gauss.thy	Sat Jun 28 09:16:42 2014 +0200
    35.3 @@ -239,12 +239,13 @@
    35.4    by (auto simp add: C_eq card_Un_disjoint D_E_disj finite_D finite_E)
    35.5  
    35.6  lemma C_prod_eq_D_times_E: "setprod id E * setprod id D = setprod id C"
    35.7 -  by (metis C_eq D_E_disj finite_D finite_E inf_commute setprod_Un_disjoint sup_commute)
    35.8 +  by (metis C_eq D_E_disj finite_D finite_E inf_commute setprod.union_disjoint sup_commute)
    35.9  
   35.10  lemma C_B_zcong_prod: "[setprod id C = setprod id B] (mod p)"
   35.11    apply (auto simp add: C_def)
   35.12    apply (insert finite_B SR_B_inj)
   35.13 -  apply (frule_tac f = "\<lambda>x. x mod int p" in setprod_reindex_id [symmetric], auto)
   35.14 +  apply (drule setprod.reindex [of "\<lambda>x. x mod int p" B id])
   35.15 +  apply auto
   35.16    apply (rule cong_setprod_int)
   35.17    apply (auto simp add: cong_int_def)
   35.18    done
   35.19 @@ -291,14 +292,14 @@
   35.20    by (auto simp add: card_seteq)
   35.21  
   35.22  lemma prod_D_F_eq_prod_A: "(setprod id D) * (setprod id F) = setprod id A"
   35.23 -  by (metis F_D_disj F_Un_D_eq_A Int_commute Un_commute finite_D finite_F setprod_Un_disjoint)
   35.24 +  by (metis F_D_disj F_Un_D_eq_A Int_commute Un_commute finite_D finite_F setprod.union_disjoint)
   35.25  
   35.26  lemma prod_F_zcong: "[setprod id F = ((-1) ^ (card E)) * (setprod id E)] (mod p)"
   35.27  proof -
   35.28    have FE: "setprod id F = setprod (op - p) E"
   35.29      apply (auto simp add: F_def)
   35.30      apply (insert finite_E inj_on_pminusx_E)
   35.31 -    apply (frule setprod_reindex_id, auto)
   35.32 +    apply (drule setprod.reindex, auto)
   35.33      done
   35.34    then have "\<forall>x \<in> E. [(p-x) mod p = - x](mod p)"
   35.35      by (metis cong_int_def minus_mod_self1 mod_mod_trivial)
   35.36 @@ -324,7 +325,7 @@
   35.37    "[a ^ nat((int p - 1) div 2) = (-1) ^ (card E)] (mod p)"
   35.38  proof -
   35.39    have "[setprod id A = setprod id F * setprod id D](mod p)"
   35.40 -    by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod_cong)
   35.41 +    by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod.cong)
   35.42    then have "[setprod id A = ((-1)^(card E) * setprod id E) * setprod id D] (mod p)"
   35.43      apply (rule cong_trans_int)
   35.44      apply (metis cong_scalar_int prod_F_zcong)
   35.45 @@ -338,7 +339,7 @@
   35.46      by (simp add: B_def)
   35.47    then have "[setprod id A = ((-1)^(card E) * (setprod (\<lambda>x. x * a) A))]
   35.48      (mod p)"
   35.49 -    by (simp add:finite_A inj_on_xa_A setprod_reindex_id[symmetric] cong del:setprod_cong)
   35.50 +    by (simp add: inj_on_xa_A setprod.reindex)
   35.51    moreover have "setprod (\<lambda>x. x * a) A =
   35.52      setprod (\<lambda>x. a) A * setprod id A"
   35.53      using finite_A by (induct set: finite) auto
   35.54 @@ -360,7 +361,7 @@
   35.55      done
   35.56    then have "[setprod id A * (-1)^(card E) = setprod id A * a^(card A)](mod p)"
   35.57      apply (rule cong_trans_int)
   35.58 -    apply (simp add: aux cong del:setprod_cong)
   35.59 +    apply (simp add: aux cong del:setprod.cong)
   35.60      done
   35.61    with A_prod_relprime have "[-1 ^ card E = a ^ card A](mod p)"
   35.62      by (metis cong_mult_lcancel_int)
    36.1 --- a/src/HOL/Number_Theory/UniqueFactorization.thy	Fri Jun 27 22:08:55 2014 +0200
    36.2 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy	Sat Jun 28 09:16:42 2014 +0200
    36.3 @@ -89,7 +89,7 @@
    36.4      then have b: "set_of N = {a} Un (set_of N - {a})"
    36.5        by auto
    36.6      then show ?thesis
    36.7 -      by (subst (1) b, subst setprod_Un_disjoint, auto)
    36.8 +      by (subst (1) b, subst setprod.union_disjoint, auto)
    36.9    next
   36.10      assume "a ~: set_of N" 
   36.11      then show ?thesis by auto
   36.12 @@ -458,18 +458,18 @@
   36.13    apply (simp only: prime_factors_altdef_nat)
   36.14    apply auto
   36.15    apply (subst power_add)
   36.16 -  apply (subst setprod_timesf)
   36.17 +  apply (subst setprod.distrib)
   36.18    apply (rule arg_cong2 [where f = "\<lambda>x y. x*y"])
   36.19    apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors k Un 
   36.20        (prime_factors l - prime_factors k)")
   36.21    apply (erule ssubst)
   36.22 -  apply (subst setprod_Un_disjoint)
   36.23 +  apply (subst setprod.union_disjoint)
   36.24    apply auto
   36.25    apply (metis One_nat_def nat_mult_1_right prime_factorization_nat setprod.neutral_const)
   36.26    apply (subgoal_tac "prime_factors k Un prime_factors l = prime_factors l Un 
   36.27        (prime_factors k - prime_factors l)")
   36.28    apply (erule ssubst)
   36.29 -  apply (subst setprod_Un_disjoint)
   36.30 +  apply (subst setprod.union_disjoint)
   36.31    apply auto
   36.32    apply (subgoal_tac "(\<Prod>p\<in>prime_factors k - prime_factors l. p ^ multiplicity p l) = 
   36.33        (\<Prod>p\<in>prime_factors k - prime_factors l. 1)")
   36.34 @@ -545,11 +545,11 @@
   36.35      [where f = "%x. nat(int(nat(f x)))", 
   36.36        transferred direction: nat "op <= (0::int)"])
   36.37    apply auto
   36.38 -  apply (subst (asm) setprod_cong)
   36.39 +  apply (subst (asm) setprod.cong)
   36.40    apply (rule refl)
   36.41    apply (rule if_P)
   36.42    apply auto
   36.43 -  apply (rule setsum_cong)
   36.44 +  apply (rule setsum.cong)
   36.45    apply auto
   36.46    done
   36.47  
   36.48 @@ -566,10 +566,10 @@
   36.49    prefer 5 apply (rule refl)
   36.50    apply (rule refl)
   36.51    apply auto
   36.52 -  apply (subst setprod_mono_one_right)
   36.53 +  apply (subst setprod.mono_neutral_right)
   36.54    apply assumption
   36.55    prefer 3
   36.56 -  apply (rule setprod_cong)
   36.57 +  apply (rule setprod.cong)
   36.58    apply (rule refl)
   36.59    apply auto
   36.60  done
    37.1 --- a/src/HOL/Old_Number_Theory/Euler.thy	Fri Jun 27 22:08:55 2014 +0200
    37.2 +++ b/src/HOL/Old_Number_Theory/Euler.thy	Sat Jun 28 09:16:42 2014 +0200
    37.3 @@ -180,7 +180,7 @@
    37.4  proof -
    37.5    from assms have "[\<Prod>(Union (SetS a p)) = setprod (setprod (%x. x)) (SetS a p)] (mod p)"
    37.6      by (auto simp add: SetS_finite SetS_elems_finite
    37.7 -      MultInvPair_prop1c setprod_Union_disjoint)
    37.8 +      MultInvPair_prop1c setprod.Union_disjoint)
    37.9    also have "[setprod (setprod (%x. x)) (SetS a p) = 
   37.10        setprod (%x. a) (SetS a p)] (mod p)"
   37.11      by (rule setprod_same_function_zcong)
    38.1 --- a/src/HOL/Old_Number_Theory/EulerFermat.thy	Fri Jun 27 22:08:55 2014 +0200
    38.2 +++ b/src/HOL/Old_Number_Theory/EulerFermat.thy	Sat Jun 28 09:16:42 2014 +0200
    38.3 @@ -252,7 +252,7 @@
    38.4     apply (simplesubst BnorRset.simps)  --{*multiple redexes*}
    38.5     apply (unfold Let_def, auto)
    38.6    apply (simp add: Bnor_fin Bnor_mem_zle_swap)
    38.7 -  apply (subst setprod_insert)
    38.8 +  apply (subst setprod.insert)
    38.9      apply (rule_tac [2] Bnor_prod_power_aux)
   38.10       apply (unfold inj_on_def)
   38.11       apply (simp_all add: mult_ac Bnor_fin Bnor_mem_zle_swap)
    39.1 --- a/src/HOL/Old_Number_Theory/Gauss.thy	Fri Jun 27 22:08:55 2014 +0200
    39.2 +++ b/src/HOL/Old_Number_Theory/Gauss.thy	Sat Jun 28 09:16:42 2014 +0200
    39.3 @@ -299,14 +299,15 @@
    39.4  
    39.5  lemma C_prod_eq_D_times_E: "setprod id E * setprod id D = setprod id C"
    39.6    apply (insert D_E_disj finite_D finite_E C_eq)
    39.7 -  apply (frule setprod_Un_disjoint [of D E id])
    39.8 +  apply (frule setprod.union_disjoint [of D E id])
    39.9    apply auto
   39.10    done
   39.11  
   39.12  lemma C_B_zcong_prod: "[setprod id C = setprod id B] (mod p)"
   39.13    apply (auto simp add: C_def)
   39.14    apply (insert finite_B SR_B_inj)
   39.15 -  apply (frule_tac f = "StandardRes p" in setprod_reindex_id [symmetric], auto)
   39.16 +  apply (frule setprod.reindex [of "StandardRes p" B id])
   39.17 +  apply auto
   39.18    apply (rule setprod_same_function_zcong)
   39.19    apply (auto simp add: StandardRes_prop1 zcong_sym p_g_0)
   39.20    done
   39.21 @@ -378,7 +379,7 @@
   39.22  lemma prod_D_F_eq_prod_A:
   39.23      "(setprod id D) * (setprod id F) = setprod id A"
   39.24    apply (insert F_D_disj finite_D finite_F)
   39.25 -  apply (frule setprod_Un_disjoint [of F D id])
   39.26 +  apply (frule setprod.union_disjoint [of F D id])
   39.27    apply (auto simp add: F_Un_D_eq_A)
   39.28    done
   39.29  
   39.30 @@ -390,7 +391,8 @@
   39.31    then have "setprod id F = setprod (op - p) E"
   39.32      apply simp
   39.33      apply (insert finite_E inj_on_pminusx_E)
   39.34 -    apply (frule_tac f = "op - p" in setprod_reindex_id, auto)
   39.35 +    apply (frule setprod.reindex [of "minus p" E id])
   39.36 +    apply auto
   39.37      done
   39.38    then have one:
   39.39      "[setprod id F = setprod (StandardRes p o (op - p)) E] (mod p)"
   39.40 @@ -441,11 +443,11 @@
   39.41    "[a ^ nat((p - 1) div 2) = (-1) ^ (card E)] (mod p)"
   39.42  proof -
   39.43    have "[setprod id A = setprod id F * setprod id D](mod p)"
   39.44 -    by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod_cong)
   39.45 +    by (auto simp add: prod_D_F_eq_prod_A mult_commute cong del:setprod.cong)
   39.46    then have "[setprod id A = ((-1)^(card E) * setprod id E) *
   39.47        setprod id D] (mod p)"
   39.48      apply (rule zcong_trans)
   39.49 -    apply (auto simp add: prod_F_zcong zcong_scalar cong del: setprod_cong)
   39.50 +    apply (auto simp add: prod_F_zcong zcong_scalar cong del: setprod.cong)
   39.51      done
   39.52    then have "[setprod id A = ((-1)^(card E) * setprod id C)] (mod p)"
   39.53      apply (rule zcong_trans)
   39.54 @@ -454,14 +456,14 @@
   39.55      done
   39.56    then have "[setprod id A = ((-1)^(card E) * setprod id B)] (mod p)"
   39.57      apply (rule zcong_trans)
   39.58 -    apply (simp add: C_B_zcong_prod zcong_scalar2 cong del:setprod_cong)
   39.59 +    apply (simp add: C_B_zcong_prod zcong_scalar2 cong del:setprod.cong)
   39.60      done
   39.61    then have "[setprod id A = ((-1)^(card E) *
   39.62      (setprod id ((%x. x * a) ` A)))] (mod p)"
   39.63      by (simp add: B_def)
   39.64    then have "[setprod id A = ((-1)^(card E) * (setprod (%x. x * a) A))]
   39.65      (mod p)"
   39.66 -    by (simp add:finite_A inj_on_xa_A setprod_reindex_id[symmetric] cong del:setprod_cong)
   39.67 +    by (simp add:finite_A inj_on_xa_A setprod.reindex cong del:setprod.cong)
   39.68    moreover have "setprod (%x. x * a) A =
   39.69      setprod (%x. a) A * setprod id A"
   39.70      using finite_A by (induct set: finite) auto
   39.71 @@ -484,7 +486,7 @@
   39.72    then have "[setprod id A * (-1)^(card E) = setprod id A *
   39.73        a^(card A)](mod p)"
   39.74      apply (rule zcong_trans)
   39.75 -    apply (simp add: aux cong del:setprod_cong)
   39.76 +    apply (simp add: aux cong del:setprod.cong)
   39.77      done
   39.78    with this zcong_cancel2 [of p "setprod id A" "-1 ^ card E" "a ^ card A"]
   39.79        p_g_0 A_prod_relprime have "[-1 ^ card E = a ^ card A](mod p)"
    40.1 --- a/src/HOL/Old_Number_Theory/Pocklington.thy	Fri Jun 27 22:08:55 2014 +0200
    40.2 +++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Sat Jun 28 09:16:42 2014 +0200
    40.3 @@ -575,7 +575,7 @@
    40.4  lemma nproduct_cmul:
    40.5    assumes fS:"finite S"
    40.6    shows "setprod (\<lambda>m. (c::'a::{comm_monoid_mult})* a(m)) S = c ^ (card S) * setprod a S"
    40.7 -unfolding setprod_timesf setprod_constant[OF fS, of c] ..
    40.8 +unfolding setprod.distrib setprod_constant[OF fS, of c] ..
    40.9  
   40.10  lemma coprime_nproduct:
   40.11    assumes fS: "finite S" and Sn: "\<forall>x\<in>S. coprime n (a x)"
    41.1 --- a/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy	Fri Jun 27 22:08:55 2014 +0200
    41.2 +++ b/src/HOL/Old_Number_Theory/Quadratic_Reciprocity.thy	Sat Jun 28 09:16:42 2014 +0200
    41.3 @@ -24,20 +24,20 @@
    41.4    also have "setsum (%x. a * x) = setsum (%x. x * a)"
    41.5      by (auto simp add: mult_commute)
    41.6    also have "setsum (%x. x * a) A = setsum id B"
    41.7 -    by (simp add: B_def setsum_reindex_id[OF inj_on_xa_A])
    41.8 +    by (simp add: B_def setsum.reindex [OF inj_on_xa_A])
    41.9    also have "... = setsum (%x. p * (x div p) + StandardRes p x) B"
   41.10      by (auto simp add: StandardRes_def zmod_zdiv_equality)
   41.11    also have "... = setsum (%x. p * (x div p)) B + setsum (StandardRes p) B"
   41.12 -    by (rule setsum_addf)
   41.13 +    by (rule setsum.distrib)
   41.14    also have "setsum (StandardRes p) B = setsum id C"
   41.15 -    by (auto simp add: C_def setsum_reindex_id[OF SR_B_inj])
   41.16 +    by (auto simp add: C_def setsum.reindex [OF SR_B_inj])
   41.17    also from C_eq have "... = setsum id (D \<union> E)"
   41.18      by auto
   41.19    also from finite_D finite_E have "... = setsum id D + setsum id E"
   41.20 -    by (rule setsum_Un_disjoint) (auto simp add: D_def E_def)
   41.21 +    by (rule setsum.union_disjoint) (auto simp add: D_def E_def)
   41.22    also have "setsum (%x. p * (x div p)) B =
   41.23        setsum ((%x. p * (x div p)) o (%x. (x * a))) A"
   41.24 -    by (auto simp add: B_def setsum_reindex inj_on_xa_A)
   41.25 +    by (auto simp add: B_def setsum.reindex inj_on_xa_A)
   41.26    also have "... = setsum (%x. p * ((x * a) div p)) A"
   41.27      by (auto simp add: o_def)
   41.28    also from finite_A have "setsum (%x. p * ((x * a) div p)) A =
   41.29 @@ -53,12 +53,12 @@
   41.30      by (simp add: Un_commute)
   41.31    also from F_D_disj finite_D finite_F
   41.32    have "... = setsum id D + setsum id F"
   41.33 -    by (auto simp add: Int_commute intro: setsum_Un_disjoint)
   41.34 +    by (auto simp add: Int_commute intro: setsum.union_disjoint)
   41.35    also from F_def have "F = (%x. (p - x)) ` E"
   41.36      by auto
   41.37    also from finite_E inj_on_pminusx_E have "setsum id ((%x. (p - x)) ` E) =
   41.38        setsum (%x. (p - x)) E"
   41.39 -    by (auto simp add: setsum_reindex)
   41.40 +    by (auto simp add: setsum.reindex)
   41.41    also from finite_E have "setsum (op - p) E = setsum (%x. p) E - setsum id E"
   41.42      by (auto simp add: setsum_subtractf id_def)
   41.43    also from finite_E have "setsum (%x. p) E = p * int(card E)"
   41.44 @@ -527,7 +527,7 @@
   41.45    ultimately have "int(card (S1)) = setsum (%j. int(card (f1 j))) P_set"
   41.46      by auto
   41.47    also have "... = setsum (%j. q * j div p) P_set"
   41.48 -    using aux3a by(fastforce intro: setsum_cong)
   41.49 +    using aux3a by(fastforce intro: setsum.cong)
   41.50    finally show ?thesis .
   41.51  qed
   41.52  
   41.53 @@ -551,7 +551,7 @@
   41.54    ultimately have "int(card (S2)) = setsum (%j. int(card (f2 j))) Q_set"
   41.55      by auto
   41.56    also have "... = setsum (%j. p * j div q) Q_set"
   41.57 -    using aux3b by(fastforce intro: setsum_cong)
   41.58 +    using aux3b by(fastforce intro: setsum.cong)
   41.59    finally show ?thesis .
   41.60  qed
   41.61  
    42.1 --- a/src/HOL/Old_Number_Theory/WilsonBij.thy	Fri Jun 27 22:08:55 2014 +0200
    42.2 +++ b/src/HOL/Old_Number_Theory/WilsonBij.thy	Sat Jun 28 09:16:42 2014 +0200
    42.3 @@ -229,9 +229,9 @@
    42.4      apply (subgoal_tac [2] "a = 1 \<or> a = p - 1")
    42.5       apply (rule_tac [3] zcong_square_zless)
    42.6          apply auto
    42.7 -  apply (subst setprod_insert)
    42.8 +  apply (subst setprod.insert)
    42.9      prefer 3
   42.10 -    apply (subst setprod_insert)
   42.11 +    apply (subst setprod.insert)
   42.12        apply (auto simp add: fin_bijER)
   42.13    apply (subgoal_tac "zcong ((a * b) * \<Prod>A) (1 * 1) p")
   42.14     apply (simp add: mult_assoc)
    43.1 --- a/src/HOL/Old_Number_Theory/WilsonRuss.thy	Fri Jun 27 22:08:55 2014 +0200
    43.2 +++ b/src/HOL/Old_Number_Theory/WilsonRuss.thy	Sat Jun 28 09:16:42 2014 +0200
    43.3 @@ -251,8 +251,8 @@
    43.4     prefer 2
    43.5     apply (subst wset.simps)
    43.6     apply (auto, unfold Let_def, auto)
    43.7 -  apply (subst setprod_insert)
    43.8 -    apply (tactic {* stac @{thm setprod_insert} 3 *})
    43.9 +  apply (subst setprod.insert)
   43.10 +    apply (tactic {* stac @{thm setprod.insert} 3 *})
   43.11        apply (subgoal_tac [5]
   43.12          "zcong (a * inv p a * (\<Prod>x\<in>wset (a - 1) p. x)) (1 * 1) p")
   43.13         prefer 5
    44.1 --- a/src/HOL/Power.thy	Fri Jun 27 22:08:55 2014 +0200
    44.2 +++ b/src/HOL/Power.thy	Sat Jun 28 09:16:42 2014 +0200
    44.3 @@ -761,6 +761,22 @@
    44.4      done
    44.5  qed
    44.6  
    44.7 +
    44.8 +subsubsection {* Generalized sum over a set *}
    44.9 +
   44.10 +lemma setsum_zero_power [simp]:
   44.11 +  fixes c :: "nat \<Rightarrow> 'a::division_ring"
   44.12 +  shows "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"
   44.13 +apply (cases "finite A")
   44.14 +  by (induction A rule: finite_induct) auto
   44.15 +
   44.16 +lemma setsum_zero_power' [simp]:
   44.17 +  fixes c :: "nat \<Rightarrow> 'a::field"
   44.18 +  shows "(\<Sum>i\<in>A. c i * 0^i / d i) = (if finite A \<and> 0 \<in> A then c 0 / d 0 else 0)"
   44.19 +  using setsum_zero_power [of "\<lambda>i. c i / d i" A]
   44.20 +  by auto
   44.21 +
   44.22 +
   44.23  subsubsection {* Generalized product over a set *}
   44.24  
   44.25  lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
   44.26 @@ -768,6 +784,17 @@
   44.27  apply auto
   44.28  done
   44.29  
   44.30 +lemma setprod_power_distrib:
   44.31 +  fixes f :: "'a \<Rightarrow> 'b::comm_semiring_1"
   44.32 +  shows "setprod f A ^ n = setprod (\<lambda>x. (f x) ^ n) A"
   44.33 +proof (cases "finite A") 
   44.34 +  case True then show ?thesis 
   44.35 +    by (induct A rule: finite_induct) (auto simp add: power_mult_distrib)
   44.36 +next
   44.37 +  case False then show ?thesis 
   44.38 +    by simp
   44.39 +qed
   44.40 +
   44.41  lemma setprod_gen_delta:
   44.42    assumes fS: "finite S"
   44.43    shows "setprod (\<lambda>k. if k=a then b k else c) S = (if a \<in> S then (b a ::'a::comm_monoid_mult) * c^ (card S - 1) else c^ card S)"
   44.44 @@ -784,14 +811,14 @@
   44.45      have dj: "?A \<inter> ?B = {}" by simp
   44.46      from fS have fAB: "finite ?A" "finite ?B" by auto  
   44.47      have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
   44.48 -      apply (rule setprod_cong) by auto
   44.49 +      apply (rule setprod.cong) by auto
   44.50      have cA: "card ?A = card S - 1" using fS a by auto
   44.51      have fA1: "setprod ?f ?A = c ^ card ?A"  unfolding fA0 apply (rule setprod_constant) using fS by auto
   44.52      have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
   44.53 -      using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
   44.54 +      using setprod.union_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
   44.55        by simp
   44.56      then have ?thesis using a cA
   44.57 -      by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)}
   44.58 +      by (simp add: fA1 field_simps cong add: setprod.cong cong del: if_weak_cong)}
   44.59    ultimately show ?thesis by blast
   44.60  qed
   44.61  
    45.1 --- a/src/HOL/Probability/Bochner_Integration.thy	Fri Jun 27 22:08:55 2014 +0200
    45.2 +++ b/src/HOL/Probability/Bochner_Integration.thy	Sat Jun 28 09:16:42 2014 +0200
    45.3 @@ -165,7 +165,7 @@
    45.4    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)"
    45.5    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)"
    45.6    unfolding indicator_def
    45.7 -  using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm)
    45.8 +  using assms by (auto intro!: setsum.mono_neutral_cong_right split: split_if_asm)
    45.9  
   45.10  lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
   45.11    fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool"
   45.12 @@ -343,7 +343,7 @@
   45.13      (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. 
   45.14        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)"
   45.15      unfolding simple_bochner_integral_def
   45.16 -  proof (safe intro!: setsum_cong scaleR_cong_right)
   45.17 +  proof (safe intro!: setsum.cong scaleR_cong_right)
   45.18      fix y assume y: "y \<in> space M" "f y \<noteq> 0"
   45.19      have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = 
   45.20          {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
   45.21 @@ -365,17 +365,17 @@
   45.22      ultimately
   45.23      show "measure M {x \<in> space M. f x = f y} =
   45.24        (\<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)"
   45.25 -      apply (simp add: setsum_cases eq)
   45.26 +      apply (simp add: setsum.If_cases eq)
   45.27        apply (subst measure_finite_Union[symmetric])
   45.28        apply (auto simp: disjoint_family_on_def)
   45.29        done
   45.30    qed
   45.31    also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. 
   45.32        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))"
   45.33 -    by (auto intro!: setsum_cong simp: scaleR_setsum_left)
   45.34 +    by (auto intro!: setsum.cong simp: scaleR_setsum_left)
   45.35    also have "\<dots> = ?r"
   45.36 -    by (subst setsum_commute)
   45.37 -       (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq)
   45.38 +    by (subst setsum.commute)
   45.39 +       (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
   45.40    finally show "simple_bochner_integral M f = ?r" .
   45.41  qed
   45.42  
   45.43 @@ -397,7 +397,7 @@
   45.44      by (intro simple_bochner_integral_partition)
   45.45         (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases)
   45.46    ultimately show ?thesis
   45.47 -    by (simp add: setsum_addf[symmetric] scaleR_add_right)
   45.48 +    by (simp add: setsum.distrib[symmetric] scaleR_add_right)
   45.49  qed
   45.50  
   45.51  lemma (in linear) simple_bochner_integral_linear:
   45.52 @@ -474,7 +474,7 @@
   45.53      unfolding simple_integral_def
   45.54      by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ereal (f x)" and v=real])
   45.55         (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
   45.56 -             intro!: setsum_cong ereal_cong_mult
   45.57 +             intro!: setsum.cong ereal_cong_mult
   45.58               simp: setsum_ereal[symmetric] times_ereal.simps(1)[symmetric] mult_ac
   45.59               simp del: setsum_ereal times_ereal.simps(1))
   45.60    also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
   45.61 @@ -879,7 +879,7 @@
   45.62      by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2
   45.63                     simple_bochner_integral_def Collect_restrict
   45.64               split: split_indicator split_indicator_asm
   45.65 -             intro!: setsum_mono_zero_cong_left arg_cong2[where f=measure])
   45.66 +             intro!: setsum.mono_neutral_cong_left arg_cong2[where f=measure])
   45.67  qed
   45.68  
   45.69  inductive integrable for M f where
   45.70 @@ -2025,19 +2025,19 @@
   45.71    shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"
   45.72  proof -
   45.73    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)"
   45.74 -    by (intro setsum_mono_zero_cong_left) auto
   45.75 +    by (intro setsum.mono_neutral_cong_left) auto
   45.76      
   45.77    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)"
   45.78      by (intro integral_cong refl) (simp add: f eq)
   45.79    also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. measure (count_space A) {a} *\<^sub>R f a)"
   45.80 -    by (subst integral_setsum) (auto intro!: setsum_cong)
   45.81 +    by (subst integral_setsum) (auto intro!: setsum.cong)
   45.82    finally show ?thesis
   45.83      by auto
   45.84  qed
   45.85  
   45.86  lemma lebesgue_integral_count_space_finite: "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
   45.87    by (subst lebesgue_integral_count_space_finite_support)
   45.88 -     (auto intro!: setsum_mono_zero_cong_left)
   45.89 +     (auto intro!: setsum.mono_neutral_cong_left)
   45.90  
   45.91  subsection {* Point measure *}
   45.92  
   45.93 @@ -2386,7 +2386,7 @@
   45.94        (\<Sum>z\<in>s i ` (space N \<times> space M). measure M {y \<in> space M. s i (x, y) = z} *\<^sub>R z)"
   45.95        using s(1)[THEN simple_functionD(1)]
   45.96        unfolding simple_bochner_integral_def
   45.97 -      by (intro setsum_mono_zero_cong_left)
   45.98 +      by (intro setsum.mono_neutral_cong_left)
   45.99           (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) }
  45.100    note eq = this
  45.101  
  45.102 @@ -2752,7 +2752,7 @@
  45.103      by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset)
  45.104    interpret I: finite_product_sigma_finite M I by default fact
  45.105    have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
  45.106 -    using `i \<notin> I` by (auto intro!: setprod_cong)
  45.107 +    using `i \<notin> I` by (auto intro!: setprod.cong)
  45.108    show ?case
  45.109      unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]]
  45.110      by (simp add: * insert prod subset_insertI)
    46.1 --- a/src/HOL/Probability/Caratheodory.thy	Fri Jun 27 22:08:55 2014 +0200
    46.2 +++ b/src/HOL/Probability/Caratheodory.thy	Sat Jun 28 09:16:42 2014 +0200
    46.3 @@ -22,7 +22,7 @@
    46.4    have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
    46.5      using assms by (simp add: fun_eq_iff)
    46.6    have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
    46.7 -    by (simp add: setsum_reindex[OF inj_prod_decode] comp_def)
    46.8 +    by (simp add: setsum.reindex[OF inj_prod_decode] comp_def)
    46.9    { fix n
   46.10      let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
   46.11      { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
   46.12 @@ -40,7 +40,7 @@
   46.13        by (auto intro!: setsum_mono3 simp: pos) }
   46.14    ultimately
   46.15    show ?thesis unfolding g_def using pos
   46.16 -    by (auto intro!: SUP_eq  simp: setsum_cartesian_product reindex SUP_upper2
   46.17 +    by (auto intro!: SUP_eq  simp: setsum.cartesian_product reindex SUP_upper2
   46.18                       setsum_nonneg suminf_ereal_eq_SUP SUP_pair
   46.19                       SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
   46.20  qed
   46.21 @@ -712,7 +712,7 @@
   46.22    with `volume M f` have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
   46.23      unfolding volume_def by blast
   46.24    also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
   46.25 -  proof (subst setsum_reindex_nonzero)
   46.26 +  proof (subst setsum.reindex_nontrivial)
   46.27      fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j"
   46.28      with `disjoint_family_on A I` have "A i = {}"
   46.29        by (auto simp: disjoint_family_on_def)
   46.30 @@ -753,7 +753,7 @@
   46.31      fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
   46.32      assume "\<Union>C = \<Union>D"
   46.33      have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))"
   46.34 -    proof (intro setsum_cong refl)
   46.35 +    proof (intro setsum.cong refl)
   46.36        fix d assume "d \<in> D"
   46.37        have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d"
   46.38          using `d \<in> D` `\<Union>C = \<Union>D` by auto
   46.39 @@ -775,7 +775,7 @@
   46.40      assume "\<Union>C = \<Union>D"
   46.41      with split_sum[OF C D] split_sum[OF D C]
   46.42      have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
   46.43 -      by (simp, subst setsum_commute, simp add: ac_simps) }
   46.44 +      by (simp, subst setsum.commute, simp add: ac_simps) }
   46.45    note sum_eq = this
   46.46  
   46.47    { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
   46.48 @@ -810,7 +810,7 @@
   46.49      also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)"
   46.50        using C_Int_cases volume_empty[OF `volume M \<mu>`] by (elim disjE) simp_all
   46.51      also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
   46.52 -      using Ca Cb by (simp add: setsum_Un_Int)
   46.53 +      using Ca Cb by (simp add: setsum.union_inter)
   46.54      also have "\<dots> = \<mu>' a + \<mu>' b"
   46.55        using Ca Cb by (simp add: \<mu>')
   46.56      finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b"
   46.57 @@ -863,7 +863,7 @@
   46.58          by (simp add: sums_iff)
   46.59      qed
   46.60      also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)"
   46.61 -      using F'(2) by (subst (2) F') (simp add: setsum_reindex)
   46.62 +      using F'(2) by (subst (2) F') (simp add: setsum.reindex)
   46.63      finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" .
   46.64    next
   46.65      show "\<mu> {} = 0"
   46.66 @@ -983,7 +983,7 @@
   46.67        by (intro suminf_setsum_ereal positiveD2[OF pos] G.Int G.finite_Union)
   46.68           (auto intro: generated_ringI_Basic)
   46.69      also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
   46.70 -      using eq V C' by (auto intro!: setsum_cong)
   46.71 +      using eq V C' by (auto intro!: setsum.cong)
   46.72      also have "\<dots> = \<mu>_r (\<Union>C')"
   46.73        using C' Un_A
   46.74        by (subst volume_finite_additive[symmetric, OF V(1)])
    47.1 --- a/src/HOL/Probability/Distributions.thy	Fri Jun 27 22:08:55 2014 +0200
    47.2 +++ b/src/HOL/Probability/Distributions.thy	Sat Jun 28 09:16:42 2014 +0200
    47.3 @@ -126,7 +126,7 @@
    47.4      qed
    47.5    qed auto
    47.6    also have "\<dots> = ereal (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n))"
    47.7 -    by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] setsum_cases)
    47.8 +    by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] setsum.If_cases)
    47.9    finally show ?thesis
   47.10      by (cases "?I") (auto simp: field_simps)
   47.11  qed
    48.1 --- a/src/HOL/Probability/Finite_Product_Measure.thy	Fri Jun 27 22:08:55 2014 +0200
    48.2 +++ b/src/HOL/Probability/Finite_Product_Measure.thy	Sat Jun 28 09:16:42 2014 +0200
    48.3 @@ -636,12 +636,12 @@
    48.4        by (auto simp: prod_emb_iff PiE_def Pi_iff split: split_if_asm) blast+
    48.5      also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
    48.6        (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
    48.7 -      using E by (subst insert) (auto intro!: setprod_cong)
    48.8 +      using E by (subst insert) (auto intro!: setprod.cong)
    48.9      also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
   48.10         emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
   48.11 -      using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod_cong)
   48.12 +      using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod.cong)
   48.13      also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
   48.14 -      using insert(1,2) J E by (intro setprod_mono_one_right) auto
   48.15 +      using insert(1,2) J E by (intro setprod.mono_neutral_right) auto
   48.16      finally show "?\<mu> ?p = \<dots>" .
   48.17  
   48.18      show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))"
   48.19 @@ -653,7 +653,7 @@
   48.20      show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
   48.21        insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))"
   48.22        using insert by auto
   48.23 -  qed (auto intro!: setprod_cong)
   48.24 +  qed (auto intro!: setprod.cong)
   48.25    with insert show ?case
   48.26      by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
   48.27  qed simp
   48.28 @@ -764,7 +764,7 @@
   48.29        using `finite J` `finite I` F unfolding X
   48.30        by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times)
   48.31      also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
   48.32 -      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
   48.33 +      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod.union_inter_neutral)
   48.34      also have "\<dots> = emeasure ?P (Pi\<^sub>E (I \<union> J) F)"
   48.35        using `finite J` `finite I` F unfolding A
   48.36        by (intro IJ.measure_times[symmetric]) auto
   48.37 @@ -849,7 +849,7 @@
   48.38    note `finite I`[intro, simp]
   48.39    interpret I: finite_product_sigma_finite M I by default auto
   48.40    have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
   48.41 -    using insert by (auto intro!: setprod_cong)
   48.42 +    using insert by (auto intro!: setprod.cong)
   48.43    have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
   48.44      using sets.sets_into_space insert
   48.45      by (intro borel_measurable_ereal_setprod
    49.1 --- a/src/HOL/Probability/Independent_Family.thy	Fri Jun 27 22:08:55 2014 +0200
    49.2 +++ b/src/HOL/Probability/Independent_Family.thy	Sat Jun 28 09:16:42 2014 +0200
    49.3 @@ -203,7 +203,7 @@
    49.4              have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))"
    49.5                using J A `j \<in> K` by (intro indep_setsD[OF G']) auto
    49.6              then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))"
    49.7 -              using `finite J` `j \<notin> J` by (auto intro!: setprod_cong) }
    49.8 +              using `finite J` `j \<notin> J` by (auto intro!: setprod.cong) }
    49.9            ultimately have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = (prob (space M) - prob X) * (\<Prod>i\<in>J. prob (A i))"
   49.10              by (simp add: field_simps)
   49.11            also have "\<dots> = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))"
   49.12 @@ -229,7 +229,7 @@
   49.13            qed
   49.14            moreover { fix k
   49.15              from J A `j \<in> K` have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))"
   49.16 -              by (subst indep_setsD[OF F(2)]) (auto intro!: setprod_cong split: split_if_asm)
   49.17 +              by (subst indep_setsD[OF F(2)]) (auto intro!: setprod.cong split: split_if_asm)
   49.18              also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)"
   49.19                using J A `j \<in> K` by (subst indep_setsD[OF G(1)]) auto
   49.20              finally have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * prob (\<Inter>i\<in>J. A i)" . }
   49.21 @@ -455,9 +455,9 @@
   49.22        also have "\<dots> = (\<Prod>l\<in>(\<Union>k\<in>K. L k). prob (?E' l))"
   49.23          using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono)
   49.24        also have "\<dots> = (\<Prod>j\<in>K. \<Prod>l\<in>L j. prob (E' j l))"
   49.25 -        using K L L_inj by (subst setprod_UN_disjoint) auto
   49.26 +        using K L L_inj by (subst setprod.UNION_disjoint) auto
   49.27        also have "\<dots> = (\<Prod>j\<in>K. prob (A j))"
   49.28 -        using K L E' by (auto simp add: A intro!: setprod_cong indep_setsD[OF indep, symmetric]) blast
   49.29 +        using K L E' by (auto simp add: A intro!: setprod.cong indep_setsD[OF indep, symmetric]) blast
   49.30        finally show "prob (\<Inter>j\<in>K. A j) = (\<Prod>j\<in>K. prob (A j))" .
   49.31      qed
   49.32    next
   49.33 @@ -799,7 +799,7 @@
   49.34        by auto
   49.35      also have "\<dots> = (\<Prod>j\<in>J. prob (A j))"
   49.36        unfolding if_distrib setprod.If_cases[OF `finite I`]
   49.37 -      using prob_space `J \<subseteq> I` by (simp add: Int_absorb1 setprod_1)
   49.38 +      using prob_space `J \<subseteq> I` by (simp add: Int_absorb1 setprod.neutral_const)
   49.39      finally show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" ..
   49.40    qed
   49.41  qed
   49.42 @@ -1225,7 +1225,7 @@
   49.43      by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def)
   49.44  
   49.45    have "(\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. X i \<omega>) \<partial>M) = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (Y i \<omega>)) \<partial>M)"
   49.46 -    using I(3) by (auto intro!: nn_integral_cong setprod_cong simp add: Y_def max_def)
   49.47 +    using I(3) by (auto intro!: nn_integral_cong setprod.cong simp add: Y_def max_def)
   49.48    also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>distr M (Pi\<^sub>M I (\<lambda>i. borel)) (\<lambda>x. \<lambda>i\<in>I. Y i x))"
   49.49      by (subst nn_integral_distr) auto
   49.50    also have "\<dots> = (\<integral>\<^sup>+\<omega>. (\<Prod>i\<in>I. max 0 (\<omega> i)) \<partial>Pi\<^sub>M I (\<lambda>i. distr M borel (Y i)))"
   49.51 @@ -1233,7 +1233,7 @@
   49.52    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^sup>+\<omega>. max 0 \<omega> \<partial>distr M borel (Y i)))"
   49.53      by (rule product_nn_integral_setprod) (auto intro: `finite I`)
   49.54    also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+\<omega>. X i \<omega> \<partial>M)"
   49.55 -    by (intro setprod_cong nn_integral_cong)
   49.56 +    by (intro setprod.cong nn_integral_cong)
   49.57         (auto simp: nn_integral_distr nn_integral_max_0 Y_def rv_X)
   49.58    finally show ?thesis .
   49.59  qed (simp add: emeasure_space_1)
   49.60 @@ -1275,7 +1275,7 @@
   49.61    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<omega>. \<omega> \<partial>distr M borel (Y i)))"
   49.62      by (rule product_integral_setprod) (auto intro: `finite I` simp: integrable_distr_eq int_Y)
   49.63    also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<omega>. X i \<omega> \<partial>M)"
   49.64 -    by (intro setprod_cong integral_cong)
   49.65 +    by (intro setprod.cong integral_cong)
   49.66         (auto simp: integral_distr Y_def rv_X)
   49.67    finally show ?eq .
   49.68  
    50.1 --- a/src/HOL/Probability/Infinite_Product_Measure.thy	Fri Jun 27 22:08:55 2014 +0200
    50.2 +++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Sat Jun 28 09:16:42 2014 +0200
    50.3 @@ -322,12 +322,12 @@
    50.4        using J  `I \<noteq> {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
    50.5      also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
    50.6        if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
    50.7 -      using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
    50.8 +      using J `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
    50.9      finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = \<dots>" .
   50.10    next
   50.11      let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"
   50.12      have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)"
   50.13 -      using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
   50.14 +      using X `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
   50.15      then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
   50.16        emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
   50.17        using X by (auto simp add: emeasure_PiM) 
   50.18 @@ -560,12 +560,12 @@
   50.19    also have "emeasure S ?F = (\<Prod>j\<in>(op + i) -` J. emeasure M (E (i + j)))"
   50.20      using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI inj_on_def)
   50.21    also have "\<dots> = (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j))"
   50.22 -    by (rule strong_setprod_reindex_cong[where f="\<lambda>x. x - i"])
   50.23 +    by (rule setprod.reindex_cong [of "\<lambda>x. x - i"])
   50.24         (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)
   50.25    also have "emeasure S ?E = (\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j))"
   50.26      using J by (intro emeasure_PiM_emb) simp_all
   50.27    also have "(\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j)) * (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
   50.28 -    by (subst mult_commute) (auto simp: J setprod_subset_diff[symmetric])
   50.29 +    by (subst mult_commute) (auto simp: J setprod.subset_diff[symmetric])
   50.30    finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" .
   50.31  qed simp_all
   50.32  
   50.33 @@ -591,7 +591,7 @@
   50.34    also have "emeasure S ?F = (\<Prod>j\<in>Suc -` J. emeasure M (E (Suc j)))"
   50.35      using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI)
   50.36    also have "\<dots> = (\<Prod>j\<in>J - {0}. emeasure M (E j))"
   50.37 -    by (rule strong_setprod_reindex_cong[where f="\<lambda>x. x - 1"])
   50.38 +    by (rule setprod.reindex_cong [of "\<lambda>x. x - 1"])
   50.39         (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)
   50.40    also have "emeasure M ?E * (\<Prod>j\<in>J - {0}. emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
   50.41      by (auto simp: M.emeasure_space_1 setprod.remove J)
    51.1 --- a/src/HOL/Probability/Information.thy	Fri Jun 27 22:08:55 2014 +0200
    51.2 +++ b/src/HOL/Probability/Information.thy	Sat Jun 28 09:16:42 2014 +0200
    51.3 @@ -19,7 +19,7 @@
    51.4  
    51.5  lemma setsum_cartesian_product':
    51.6    "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
    51.7 -  unfolding setsum_cartesian_product by simp
    51.8 +  unfolding setsum.cartesian_product by simp
    51.9  
   51.10  lemma split_pairs:
   51.11    "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
   51.12 @@ -735,8 +735,8 @@
   51.13      by auto
   51.14    with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^sub>M count_space (Y ` space M))) =
   51.15      (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
   51.16 -    by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum_cases split_beta'
   51.17 -             intro!: setsum_cong)
   51.18 +    by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum.If_cases split_beta'
   51.19 +             intro!: setsum.cong)
   51.20  qed
   51.21  
   51.22  lemma (in information_space)
   51.23 @@ -748,7 +748,7 @@
   51.24  proof (subst mutual_information_simple_distributed[OF Px Py Pxy])
   51.25    have "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) =
   51.26      (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. 0)"
   51.27 -    by (intro setsum_cong) (auto simp: ae)
   51.28 +    by (intro setsum.cong) (auto simp: ae)
   51.29    then show "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M.
   51.30      Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp
   51.31  qed
   51.32 @@ -1488,7 +1488,7 @@
   51.33    have "(\<lambda>(x, y, z). ?i x y z) = (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)"
   51.34      by (auto intro!: ext)
   51.35    then show "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x)) ` space M. ?j x y z)"
   51.36 -    by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite setsum_cases split_beta')
   51.37 +    by (auto intro!: setsum.cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite setsum.If_cases split_beta')
   51.38  qed
   51.39  
   51.40  lemma (in information_space) conditional_mutual_information_nonneg:
   51.41 @@ -1640,7 +1640,7 @@
   51.42      by auto
   51.43    from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
   51.44      - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
   51.45 -    by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum_cases split_beta')
   51.46 +    by (auto intro!: setsum.cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum.If_cases split_beta')
   51.47  qed
   51.48  
   51.49  lemma (in information_space) conditional_mutual_information_eq_conditional_entropy:
   51.50 @@ -1671,7 +1671,7 @@
   51.51    then show ?thesis
   51.52      apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy])
   51.53      apply (subst conditional_entropy_eq[OF Py Pxy])
   51.54 -    apply (auto intro!: setsum_cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum_reindex[OF inj]
   51.55 +    apply (auto intro!: setsum.cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum.reindex[OF inj]
   51.56                  log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space)
   51.57      using Py[THEN simple_distributed, THEN distributed_real_AE] Pxy[THEN simple_distributed, THEN distributed_real_AE]
   51.58    apply (auto simp add: not_le[symmetric] AE_count_space)
   51.59 @@ -1857,13 +1857,13 @@
   51.60    have "\<H>(\<lambda>x. (X x, Y x)) = - (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` space M. ?f x * log b (?f x))"
   51.61      using XY by (rule entropy_simple_distributed)
   51.62    also have "\<dots> = - (\<Sum>x\<in>(\<lambda>(x, y). (y, x)) ` (\<lambda>x. (X x, Y x)) ` space M. ?g x * log b (?g x))"
   51.63 -    by (subst (2) setsum_reindex) (auto simp: inj_on_def intro!: setsum_cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
   51.64 +    by (subst (2) setsum.reindex) (auto simp: inj_on_def intro!: setsum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
   51.65    also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))"
   51.66 -    by (auto intro!: setsum_cong)
   51.67 +    by (auto intro!: setsum.cong)
   51.68    also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
   51.69      by (subst entropy_distr[OF simple_distributed_joint[OF YX]])
   51.70         (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
   51.71 -             cong del: setsum_cong  intro!: setsum_mono_zero_left)
   51.72 +             cong del: setsum.cong  intro!: setsum.mono_neutral_left)
   51.73    finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^sub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
   51.74    then show ?thesis
   51.75      unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp
   51.76 @@ -1882,8 +1882,8 @@
   51.77      apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] refl]])
   51.78      apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
   51.79      unfolding eq
   51.80 -    apply (subst setsum_reindex[OF inj])
   51.81 -    apply (auto intro!: setsum_cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
   51.82 +    apply (subst setsum.reindex[OF inj])
   51.83 +    apply (auto intro!: setsum.cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
   51.84      done
   51.85  qed
   51.86  
   51.87 @@ -1908,7 +1908,7 @@
   51.88      unfolding o_assoc
   51.89      apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
   51.90      apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\<lambda>x. prob (X -` {x} \<inter> space M)"])
   51.91 -    apply (auto intro!: setsum_cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def)
   51.92 +    apply (auto intro!: setsum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def)
   51.93      done
   51.94    also have "... \<le> \<H>(f \<circ> X)"
   51.95      using entropy_data_processing[OF sf] .
    52.1 --- a/src/HOL/Probability/Measure_Space.thy	Fri Jun 27 22:08:55 2014 +0200
    52.2 +++ b/src/HOL/Probability/Measure_Space.thy	Sat Jun 28 09:16:42 2014 +0200
    52.3 @@ -51,7 +51,7 @@
    52.4    have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
    52.5      using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
    52.6    then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
    52.7 -    by (auto simp: setsum_cases)
    52.8 +    by (auto simp: setsum.If_cases)
    52.9    moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
   52.10    proof (rule SUP_eqI)
   52.11      fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
   52.12 @@ -287,7 +287,7 @@
   52.13    from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
   52.14      by (rule suminf_finite) auto
   52.15    also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
   52.16 -    using fin_not_empty F_subset by (rule setsum_mono_zero_left) auto
   52.17 +    using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
   52.18    also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
   52.19      using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
   52.20    also have "\<dots> = \<mu> (\<Union>i. F i)"
   52.21 @@ -668,7 +668,7 @@
   52.22    then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
   52.23      using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
   52.24    also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
   52.25 -    using X eq by (auto intro!: setsum_cong)
   52.26 +    using X eq by (auto intro!: setsum.cong)
   52.27    also have "\<dots> = emeasure N X"
   52.28      using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
   52.29    finally show "emeasure M X = emeasure N X" .
    53.1 --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Fri Jun 27 22:08:55 2014 +0200
    53.2 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Sat Jun 28 09:16:42 2014 +0200
    53.3 @@ -72,9 +72,9 @@
    53.4  proof -
    53.5    have "?r = (\<Sum>y \<in> f ` space M.
    53.6      (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
    53.7 -    by (auto intro!: setsum_cong2)
    53.8 +    by (auto intro!: setsum.cong)
    53.9    also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
   53.10 -    using assms by (auto dest: simple_functionD simp: setsum_delta)
   53.11 +    using assms by (auto dest: simple_functionD simp: setsum.delta)
   53.12    also have "... = f x" using x by (auto simp: indicator_def)
   53.13    finally show ?thesis by auto
   53.14  qed
   53.15 @@ -552,7 +552,7 @@
   53.16      (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M. 
   53.17        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))"
   53.18      unfolding simple_integral_def
   53.19 -  proof (safe intro!: setsum_cong ereal_left_mult_cong)
   53.20 +  proof (safe intro!: setsum.cong ereal_left_mult_cong)
   53.21      fix y assume y: "y \<in> space M" "f y \<noteq> 0"
   53.22      have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = 
   53.23          {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
   53.24 @@ -565,17 +565,17 @@
   53.25        by (rule rev_finite_subset) auto
   53.26      then show "emeasure M (f -` {f y} \<inter> space M) =
   53.27        (\<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)"
   53.28 -      apply (simp add: setsum_cases)
   53.29 +      apply (simp add: setsum.If_cases)
   53.30        apply (subst setsum_emeasure)
   53.31        apply (auto simp: disjoint_family_on_def eq)
   53.32        done
   53.33    qed
   53.34    also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. 
   53.35        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))"
   53.36 -    by (auto intro!: setsum_cong simp: setsum_ereal_right_distrib emeasure_nonneg)
   53.37 +    by (auto intro!: setsum.cong simp: setsum_ereal_right_distrib emeasure_nonneg)
   53.38    also have "\<dots> = ?r"
   53.39 -    by (subst setsum_commute)
   53.40 -       (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq)
   53.41 +    by (subst setsum.commute)
   53.42 +       (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
   53.43    finally show "integral\<^sup>S M f = ?r" .
   53.44  qed
   53.45  
   53.46 @@ -588,7 +588,7 @@
   53.47      by (intro simple_function_partition) (auto intro: f g)
   53.48    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}) +
   53.49      (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
   53.50 -    using assms(2,4) by (auto intro!: setsum_cong ereal_left_distrib simp: setsum_addf[symmetric])
   53.51 +    using assms(2,4) by (auto intro!: setsum.cong ereal_left_distrib simp: setsum.distrib[symmetric])
   53.52    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)"
   53.53      by (intro simple_function_partition[symmetric]) (auto intro: f g)
   53.54    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)"
   53.55 @@ -686,14 +686,14 @@
   53.56      using assms by (intro simple_function_partition) auto
   53.57    also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ereal))`space M.
   53.58      if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
   53.59 -    by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum_cong)
   53.60 +    by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong)
   53.61    also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
   53.62 -    using assms by (subst setsum_cases) (auto intro!: simple_functionD(1) simp: eq)
   53.63 +    using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
   53.64    also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
   53.65 -    by (subst setsum_reindex[where f=fst]) (auto simp: inj_on_def)
   53.66 +    by (subst setsum.reindex [of fst]) (auto simp: inj_on_def)
   53.67    also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
   53.68      using A[THEN sets.sets_into_space]
   53.69 -    by (intro setsum_mono_zero_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
   53.70 +    by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
   53.71    finally show ?thesis .
   53.72  qed
   53.73  
   53.74 @@ -929,7 +929,7 @@
   53.75    next
   53.76      show "integral\<^sup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))"
   53.77        using measure_conv u_range B_u unfolding simple_integral_def
   53.78 -      by (auto intro!: setsum_cong SUP_ereal_cmult [symmetric])
   53.79 +      by (auto intro!: setsum.cong SUP_ereal_cmult [symmetric])
   53.80    qed
   53.81    moreover
   53.82    have "a * (SUP i. integral\<^sup>S M (?uB i)) \<le> ?S"
   53.83 @@ -1615,19 +1615,19 @@
   53.84    have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
   53.85      (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
   53.86      by (auto intro!: nn_integral_cong
   53.87 -             simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less)
   53.88 +             simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less)
   53.89    also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
   53.90      by (subst nn_integral_setsum)
   53.91         (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le)
   53.92    also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
   53.93 -    by (auto intro!: setsum_cong simp: nn_integral_cmult_indicator one_ereal_def[symmetric])
   53.94 +    by (auto intro!: setsum.cong simp: nn_integral_cmult_indicator one_ereal_def[symmetric])
   53.95    finally show ?thesis by (simp add: nn_integral_max_0)
   53.96  qed
   53.97  
   53.98  lemma nn_integral_count_space_finite:
   53.99      "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))"
  53.100    by (subst nn_integral_max_0[symmetric])
  53.101 -     (auto intro!: setsum_mono_zero_left simp: nn_integral_count_space less_le)
  53.102 +     (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
  53.103  
  53.104  lemma emeasure_UN_countable:
  53.105    assumes sets: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I: "countable I" 
  53.106 @@ -1636,7 +1636,7 @@
  53.107  proof cases
  53.108    assume "finite I" with sets disj show ?thesis
  53.109      by (subst setsum_emeasure[symmetric])
  53.110 -       (auto intro!: setsum_cong simp add: max_def subset_eq nn_integral_count_space_finite emeasure_nonneg)
  53.111 +       (auto intro!: setsum.cong simp add: max_def subset_eq nn_integral_count_space_finite emeasure_nonneg)
  53.112  next
  53.113    assume f: "\<not> finite I"
  53.114    then have [intro]: "I \<noteq> {}" by auto
  53.115 @@ -1786,7 +1786,7 @@
  53.116    using simple_function_restrict_space_ereal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
  53.117    by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
  53.118             split: split_indicator split_indicator_asm
  53.119 -           intro!: setsum_mono_zero_cong_left ereal_left_mult_cong arg_cong2[where f=emeasure])
  53.120 +           intro!: setsum.mono_neutral_cong_left ereal_left_mult_cong arg_cong2[where f=emeasure])
  53.121  
  53.122  lemma one_not_less_zero[simp]: "\<not> 1 < (0::ereal)"
  53.123    by (simp add: zero_ereal_def one_ereal_def) 
  53.124 @@ -2097,12 +2097,12 @@
  53.125  
  53.126  lemma emeasure_point_measure_finite:
  53.127    "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
  53.128 -  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le)
  53.129 +  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
  53.130  
  53.131  lemma emeasure_point_measure_finite2:
  53.132    "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> (\<And>i. i \<in> X \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
  53.133    by (subst emeasure_point_measure)
  53.134 -     (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le)
  53.135 +     (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
  53.136  
  53.137  lemma null_sets_point_measure_iff:
  53.138    "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x \<le> 0)"
  53.139 @@ -2121,7 +2121,7 @@
  53.140    apply (subst nn_integral_density)
  53.141    apply (simp_all add: AE_count_space nn_integral_density)
  53.142    apply (subst nn_integral_count_space )
  53.143 -  apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff)
  53.144 +  apply (auto intro!: setsum.cong simp: max_def ereal_zero_less_0_iff)
  53.145    apply (rule finite_subset)
  53.146    prefer 2
  53.147    apply assumption
  53.148 @@ -2131,7 +2131,7 @@
  53.149  lemma nn_integral_point_measure_finite:
  53.150    "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> g a) \<Longrightarrow>
  53.151      integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
  53.152 -  by (subst nn_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le)
  53.153 +  by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
  53.154  
  53.155  subsubsection {* Uniform measure *}
  53.156  
    54.1 --- a/src/HOL/Probability/Probability_Measure.thy	Fri Jun 27 22:08:55 2014 +0200
    54.2 +++ b/src/HOL/Probability/Probability_Measure.thy	Sat Jun 28 09:16:42 2014 +0200
    54.3 @@ -364,7 +364,7 @@
    54.4  sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^sub>M i\<in>I. M i"
    54.5  proof
    54.6    show "emeasure (\<Pi>\<^sub>M i\<in>I. M i) (space (\<Pi>\<^sub>M i\<in>I. M i)) = 1"
    54.7 -    by (simp add: measure_times M.emeasure_space_1 setprod_1 space_PiM)
    54.8 +    by (simp add: measure_times M.emeasure_space_1 setprod.neutral_const space_PiM)
    54.9  qed
   54.10  
   54.11  lemma (in finite_product_prob_space) prob_times:
   54.12 @@ -915,7 +915,7 @@
   54.13    from X have "setsum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)"
   54.14      by (subst finite_measure_finite_Union)
   54.15         (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD
   54.16 -             intro!: setsum_cong arg_cong[where f="prob"])
   54.17 +             intro!: setsum.cong arg_cong[where f="prob"])
   54.18    also have "\<dots> = prob (space M)"
   54.19      by (auto intro!: arg_cong[where f=prob])
   54.20    finally show ?thesis
   54.21 @@ -943,7 +943,7 @@
   54.22      Pxy[THEN simple_distributed, THEN distributed_real_AE]
   54.23    show ?thesis
   54.24      unfolding AE_count_space
   54.25 -    apply (auto simp add: nn_integral_count_space_finite * intro!: setsum_cong split: split_max)
   54.26 +    apply (auto simp add: nn_integral_count_space_finite * intro!: setsum.cong split: split_max)
   54.27      done
   54.28  qed
   54.29  
    55.1 --- a/src/HOL/Probability/Projective_Family.thy	Fri Jun 27 22:08:55 2014 +0200
    55.2 +++ b/src/HOL/Probability/Projective_Family.thy	Sat Jun 28 09:16:42 2014 +0200
    55.3 @@ -40,7 +40,7 @@
    55.4      by simp
    55.5    also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"
    55.6      using `finite K` `J \<subseteq> K`
    55.7 -    by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1)
    55.8 +    by (intro setprod.mono_neutral_left) (auto simp: M.emeasure_space_1)
    55.9    also have "\<dots> = emeasure (Pi\<^sub>M K M) (\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i))"
   55.10      using E by (simp add: K.measure_times)
   55.11    also have "(\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^sub>E J E \<inter> (\<Pi>\<^sub>E i\<in>K. space (M i))"
    56.1 --- a/src/HOL/Probability/Projective_Limit.thy	Fri Jun 27 22:08:55 2014 +0200
    56.2 +++ b/src/HOL/Probability/Projective_Limit.thy	Sat Jun 28 09:16:42 2014 +0200
    56.3 @@ -353,7 +353,7 @@
    56.4          also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps
    56.5          proof (rule mult_strict_right_mono)
    56.6            have "(\<Sum>i\<in>{1..n}. 2 powr - real i) = (\<Sum>i\<in>{1..<Suc n}. (1/2) ^ i)"
    56.7 -            by (rule setsum_cong)
    56.8 +            by (rule setsum.cong)
    56.9                 (auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide)
   56.10            also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
   56.11            also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
    57.1 --- a/src/HOL/Probability/Sigma_Algebra.thy	Fri Jun 27 22:08:55 2014 +0200
    57.2 +++ b/src/HOL/Probability/Sigma_Algebra.thy	Sat Jun 28 09:16:42 2014 +0200
    57.3 @@ -847,7 +847,7 @@
    57.4      by auto
    57.5    thus ?thesis
    57.6      unfolding indicator_def
    57.7 -    by (simp add: if_distrib setsum_cases[OF `finite P`])
    57.8 +    by (simp add: if_distrib setsum.If_cases[OF `finite P`])
    57.9  qed
   57.10  
   57.11  definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
    58.1 --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Fri Jun 27 22:08:55 2014 +0200
    58.2 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Sat Jun 28 09:16:42 2014 +0200
    58.3 @@ -99,8 +99,8 @@
    58.4      by force+
    58.5    show ?case unfolding *
    58.6      using Suc[of "\<lambda>i. P (Suc i)"]
    58.7 -    by (simp add: setsum_reindex split_conv setsum_cartesian_product'
    58.8 -      lessThan_Suc_eq_insert_0 setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
    58.9 +    by (simp add: setsum.reindex split_conv setsum_cartesian_product'
   58.10 +      lessThan_Suc_eq_insert_0 setprod.reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
   58.11  qed
   58.12  
   58.13  declare space_point_measure[simp]
   58.14 @@ -168,7 +168,7 @@
   58.15      case (Suc n)
   58.16      then show ?case
   58.17        by (simp add: lists_length_Suc_eq lessThan_Suc_eq_insert_0
   58.18 -                    setsum_reindex setprod_reindex)
   58.19 +                    setsum.reindex setprod.reindex)
   58.20    qed
   58.21    then show "setsum P msgs = 1"
   58.22      unfolding msgs_def P_def by simp
   58.23 @@ -255,7 +255,7 @@
   58.24      apply (simp add: * P_def)
   58.25      apply (simp add: setsum_cartesian_product')
   58.26      using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] `k \<in> keys`
   58.27 -    by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod_1)
   58.28 +    by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod.neutral_const)
   58.29  qed
   58.30  
   58.31  lemma fst_image_msgs[simp]: "fst`msgs = keys"
   58.32 @@ -297,8 +297,8 @@
   58.33    show "- (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {x}) * log b (\<P>(X ; Y) {x})) =
   58.34      - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` msgs. (\<P>(Y ; X) {x}) * log b (\<P>(Y ; X) {x}))"
   58.35      unfolding eq
   58.36 -    apply (subst setsum_reindex)
   58.37 -    apply (auto simp: vimage_def inj_on_def intro!: setsum_cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])
   58.38 +    apply (subst setsum.reindex)
   58.39 +    apply (auto simp: vimage_def inj_on_def intro!: setsum.cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])
   58.40      done
   58.41  qed
   58.42  
   58.43 @@ -314,15 +314,15 @@
   58.44  proof -
   58.45    have "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
   58.46      - (\<Sum>x\<in>X`msgs. (\<Sum>y\<in>Y`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
   58.47 -    unfolding setsum_cartesian_product
   58.48 -    apply (intro arg_cong[where f=uminus] setsum_mono_zero_left)
   58.49 +    unfolding setsum.cartesian_product
   58.50 +    apply (intro arg_cong[where f=uminus] setsum.mono_neutral_left)
   58.51      apply (auto simp: vimage_def image_iff intro!: measure_eq_0I)
   58.52      apply metis
   58.53      done
   58.54    also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
   58.55 -    by (subst setsum_commute) rule
   58.56 +    by (subst setsum.commute) rule
   58.57    also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
   58.58 -    by (auto simp add: setsum_right_distrib vimage_def intro!: setsum_cong prob_conj_imp1)
   58.59 +    by (auto simp add: setsum_right_distrib vimage_def intro!: setsum.cong prob_conj_imp1)
   58.60    finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
   58.61      -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
   58.62  qed
   58.63 @@ -358,7 +358,7 @@
   58.64  
   58.65      have "(\<P>(OB ; fst) {(obs, k)}) / K k = (\<P>(OB ; fst) {(obs', k)}) / K k"
   58.66        unfolding *[OF obs] *[OF obs']
   58.67 -      using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: setprod_reindex)
   58.68 +      using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: setprod.reindex)
   58.69      then have "(\<P>(OB ; fst) {(obs, k)}) = (\<P>(OB ; fst) {(obs', k)})"
   58.70        using `K k \<noteq> 0` by auto }
   58.71    note t_eq_imp = this
   58.72 @@ -396,17 +396,17 @@
   58.73        using finite_measure_mono[of "{x. t (OB x) = t obs \<and> fst x = k} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
   58.74        using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
   58.75        apply (simp add: setsum_distribution_cut[of "t\<circ>OB" "t obs" fst])
   58.76 -      apply (auto intro!: setsum_cong simp: subset_eq vimage_def prob_conj_imp1)
   58.77 +      apply (auto intro!: setsum.cong simp: subset_eq vimage_def prob_conj_imp1)
   58.78        done
   58.79      also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
   58.80        \<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
   58.81        using CP_t_K[OF `k\<in>keys` obs] CP_t_K[OF _ obs] `real (card ?S) \<noteq> 0`
   58.82        by (simp only: setsum_right_distrib[symmetric] ac_simps
   58.83            mult_divide_mult_cancel_left[OF `real (card ?S) \<noteq> 0`]
   58.84 -        cong: setsum_cong)
   58.85 +        cong: setsum.cong)
   58.86      also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
   58.87        using setsum_distribution_cut[of OB obs fst]
   58.88 -      by (auto intro!: setsum_cong simp: prob_conj_imp1 vimage_def)
   58.89 +      by (auto intro!: setsum.cong simp: prob_conj_imp1 vimage_def)
   58.90      also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
   58.91        by (auto simp: vimage_def conj_commute prob_conj_imp2)
   58.92      finally have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}" . }
   58.93 @@ -440,13 +440,13 @@
   58.94      unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp
   58.95    also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
   58.96      apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t])
   58.97 -    apply (subst setsum_reindex)
   58.98 +    apply (subst setsum.reindex)
   58.99      apply (fastforce intro!: inj_onI)
  58.100      apply simp
  58.101 -    apply (subst setsum_Sigma[symmetric, unfolded split_def])
  58.102 +    apply (subst setsum.Sigma[symmetric, unfolded split_def])
  58.103      using finite_space apply fastforce
  58.104      using finite_space apply fastforce
  58.105 -    apply (safe intro!: setsum_cong)
  58.106 +    apply (safe intro!: setsum.cong)
  58.107      using P_t_sum_P_O
  58.108      by (simp add: setsum_divide_distrib[symmetric] field_simps **
  58.109                    setsum_right_distrib[symmetric] setsum_left_distrib[symmetric])
    59.1 --- a/src/HOL/Real_Vector_Spaces.thy	Fri Jun 27 22:08:55 2014 +0200
    59.2 +++ b/src/HOL/Real_Vector_Spaces.thy	Sat Jun 28 09:16:42 2014 +0200
    59.3 @@ -450,7 +450,7 @@
    59.4      by (induct s rule: finite_induct) auto
    59.5  next
    59.6    case False then show ?thesis using assms
    59.7 -    by (metis Reals_0 setsum_infinite)
    59.8 +    by (metis Reals_0 setsum.infinite)
    59.9  qed
   59.10  
   59.11  lemma setprod_in_Reals: assumes "\<And>i. i \<in> s \<Longrightarrow> f i \<in> \<real>" shows "setprod f s \<in> \<real>"
   59.12 @@ -459,7 +459,7 @@
   59.13      by (induct s rule: finite_induct) auto
   59.14  next
   59.15    case False then show ?thesis using assms
   59.16 -    by (metis Reals_1 setprod_infinite)
   59.17 +    by (metis Reals_1 setprod.infinite)
   59.18  qed
   59.19  
   59.20  lemma Reals_induct [case_names of_real, induct set: Reals]:
    60.1 --- a/src/HOL/Series.thy	Fri Jun 27 22:08:55 2014 +0200
    60.2 +++ b/src/HOL/Series.thy	Sat Jun 28 09:16:42 2014 +0200
    60.3 @@ -75,7 +75,7 @@
    60.4      next
    60.5        assume [simp]: "N \<noteq> {}"
    60.6        show ?thesis
    60.7 -      proof (safe intro!: setsum_mono_zero_right f)
    60.8 +      proof (safe intro!: setsum.mono_neutral_right f)
    60.9          fix i assume "i \<in> N"
   60.10          then have "i \<le> Max N" by simp
   60.11          then show "i < n + Suc (Max N)" by simp
   60.12 @@ -225,7 +225,7 @@
   60.13    have "f sums (s + f 0) \<longleftrightarrow> (\<lambda>i. \<Sum>j<Suc i. f j) ----> s + f 0"
   60.14      by (subst LIMSEQ_Suc_iff) (simp add: sums_def)
   60.15    also have "\<dots> \<longleftrightarrow> (\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"
   60.16 -    by (simp add: ac_simps setsum_reindex image_iff lessThan_Suc_eq_insert_0)
   60.17 +    by (simp add: ac_simps setsum.reindex image_iff lessThan_Suc_eq_insert_0)
   60.18    also have "\<dots> \<longleftrightarrow> (\<lambda>n. f (Suc n)) sums s"
   60.19    proof
   60.20      assume "(\<lambda>i. (\<Sum>j<i. f (Suc j)) + f 0) ----> s + f 0"
   60.21 @@ -241,7 +241,7 @@
   60.22  begin
   60.23  
   60.24  lemma sums_add: "f sums a \<Longrightarrow> g sums b \<Longrightarrow> (\<lambda>n. f n + g n) sums (a + b)"
   60.25 -  unfolding sums_def by (simp add: setsum_addf tendsto_add)
   60.26 +  unfolding sums_def by (simp add: setsum.distrib tendsto_add)
   60.27  
   60.28  lemma summable_add: "summable f \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. f n + g n)"
   60.29    unfolding summable_def by (auto intro: sums_add)
   60.30 @@ -568,7 +568,7 @@
   60.31  lemma setsum_triangle_reindex:
   60.32    fixes n :: nat
   60.33    shows "(\<Sum>(i,j)\<in>{(i,j). i+j < n}. f i j) = (\<Sum>k<n. \<Sum>i\<le>k. f i (k - i))"
   60.34 -  apply (simp add: setsum_Sigma)
   60.35 +  apply (simp add: setsum.Sigma)
   60.36    apply (rule setsum.reindex_bij_witness[where j="\<lambda>(i, j). (i+j, i)" and i="\<lambda>(k, i). (i, k - i)"])
   60.37    apply auto
   60.38    done
   60.39 @@ -597,12 +597,12 @@
   60.40    have "(\<lambda>n. (\<Sum>k<n. a k) * (\<Sum>k<n. b k)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
   60.41      by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b])
   60.42    hence 1: "(\<lambda>n. setsum ?g (?S1 n)) ----> (\<Sum>k. a k) * (\<Sum>k. b k)"
   60.43 -    by (simp only: setsum_product setsum_Sigma [rule_format] finite_lessThan)
   60.44 +    by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)
   60.45  
   60.46    have "(\<lambda>n. (\<Sum>k<n. norm (a k)) * (\<Sum>k<n. norm (b k))) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
   60.47      using a b by (intro tendsto_mult summable_LIMSEQ)
   60.48    hence "(\<lambda>n. setsum ?f (?S1 n)) ----> (\<Sum>k. norm (a k)) * (\<Sum>k. norm (b k))"
   60.49 -    by (simp only: setsum_product setsum_Sigma [rule_format] finite_lessThan)
   60.50 +    by (simp only: setsum_product setsum.Sigma [rule_format] finite_lessThan)
   60.51    hence "convergent (\<lambda>n. setsum ?f (?S1 n))"
   60.52      by (rule convergentI)
   60.53    hence Cauchy: "Cauchy (\<lambda>n. setsum ?f (?S1 n))"
    61.1 --- a/src/HOL/Set_Interval.thy	Fri Jun 27 22:08:55 2014 +0200
    61.2 +++ b/src/HOL/Set_Interval.thy	Sat Jun 28 09:16:42 2014 +0200
    61.3 @@ -1378,14 +1378,14 @@
    61.4  special form for @{term"{..<n}"}. *}
    61.5  
    61.6  text{* This congruence rule should be used for sums over intervals as
    61.7 -the standard theorem @{text[source]setsum_cong} does not work well
    61.8 +the standard theorem @{text[source]setsum.cong} does not work well
    61.9  with the simplifier who adds the unsimplified premise @{term"x:B"} to
   61.10  the context. *}
   61.11  
   61.12  lemma setsum_ivl_cong:
   61.13   "\<lbrakk>a = c; b = d; !!x. \<lbrakk> c \<le> x; x < d \<rbrakk> \<Longrightarrow> f x = g x \<rbrakk> \<Longrightarrow>
   61.14   setsum f {a..<b} = setsum g {c..<d}"
   61.15 -by(rule setsum_cong, simp_all)
   61.16 +by(rule setsum.cong, simp_all)
   61.17  
   61.18  (* FIXME why are the following simp rules but the corresponding eqns
   61.19  on intervals are not? *)
   61.20 @@ -1437,13 +1437,13 @@
   61.21    shows "setsum f {m..n + p} = setsum f {m..n} + setsum f {n + 1..n + p}"
   61.22  proof-
   61.23    have "{m .. n+p} = {m..n} \<union> {n+1..n+p}" using `m \<le> n+1` by auto
   61.24 -  thus ?thesis by (auto simp: ivl_disj_int setsum_Un_disjoint
   61.25 +  thus ?thesis by (auto simp: ivl_disj_int setsum.union_disjoint
   61.26      atLeastSucAtMost_greaterThanAtMost)
   61.27  qed
   61.28  
   61.29  lemma setsum_add_nat_ivl: "\<lbrakk> m \<le> n; n \<le> p \<rbrakk> \<Longrightarrow>
   61.30    setsum f {m..<n} + setsum f {n..<p} = setsum f {m..<p::nat}"
   61.31 -by (simp add:setsum_Un_disjoint[symmetric] ivl_disj_int ivl_disj_un)
   61.32 +by (simp add:setsum.union_disjoint[symmetric] ivl_disj_int ivl_disj_un)
   61.33  
   61.34  lemma setsum_diff_nat_ivl:
   61.35  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
   61.36 @@ -1459,70 +1459,6 @@
   61.37            (if m <= n then f m - f(n + 1) else 0)"
   61.38  by (induct n, auto simp add: algebra_simps not_le le_Suc_eq)
   61.39  
   61.40 -lemma setsum_restrict_set':
   61.41 -  "finite A \<Longrightarrow> setsum f {x \<in> A. x \<in> B} = (\<Sum>x\<in>A. if x \<in> B then f x else 0)"
   61.42 -  by (simp add: setsum_restrict_set [symmetric] Int_def)
   61.43 -
   61.44 -lemma setsum_restrict_set'':
   61.45 -  "finite A \<Longrightarrow> setsum f {x \<in> A. P x} = (\<Sum>x\<in>A. if P x  then f x else 0)"
   61.46 -  by (simp add: setsum_restrict_set' [of A f "{x. P x}", simplified mem_Collect_eq])
   61.47 -
   61.48 -lemma setsum_setsum_restrict:
   61.49 -  "finite S \<Longrightarrow> finite T \<Longrightarrow>
   61.50 -    setsum (\<lambda>x. setsum (\<lambda>y. f x y) {y. y \<in> T \<and> R x y}) S = setsum (\<lambda>y. setsum (\<lambda>x. f x y) {x. x \<in> S \<and> R x y}) T"
   61.51 -  by (simp add: setsum_restrict_set'') (rule setsum_commute)
   61.52 -
   61.53 -lemma setsum_image_gen: assumes fS: "finite S"
   61.54 -  shows "setsum g S = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
   61.55 -proof-
   61.56 -  { fix x assume "x \<in> S" then have "{y. y\<in> f`S \<and> f x = y} = {f x}" by auto }
   61.57 -  hence "setsum g S = setsum (\<lambda>x. setsum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
   61.58 -    by simp
   61.59 -  also have "\<dots> = setsum (\<lambda>y. setsum g {x. x \<in> S \<and> f x = y}) (f ` S)"
   61.60 -    by (rule setsum_setsum_restrict[OF fS finite_imageI[OF fS]])
   61.61 -  finally show ?thesis .
   61.62 -qed
   61.63 -
   61.64 -lemma setsum_le_included:
   61.65 -  fixes f :: "'a \<Rightarrow> 'b::ordered_comm_monoid_add"
   61.66 -  assumes "finite s" "finite t"
   61.67 -  and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"
   61.68 -  shows "setsum f s \<le> setsum g t"
   61.69 -proof -
   61.70 -  have "setsum f s \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) s"
   61.71 -  proof (rule setsum_mono)
   61.72 -    fix y assume "y \<in> s"
   61.73 -    with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto
   61.74 -    with assms show "f y \<le> setsum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
   61.75 -      using order_trans[of "?A (i z)" "setsum g {z}" "?B (i z)", intro]
   61.76 -      by (auto intro!: setsum_mono2)
   61.77 -  qed
   61.78 -  also have "... \<le> setsum (\<lambda>y. setsum g {x. x\<in>t \<and> i x = y}) (i ` t)"
   61.79 -    using assms(2-4) by (auto intro!: setsum_mono2 setsum_nonneg)
   61.80 -  also have "... \<le> setsum g t"
   61.81 -    using assms by (auto simp: setsum_image_gen[symmetric])
   61.82 -  finally show ?thesis .
   61.83 -qed
   61.84 -
   61.85 -lemma setsum_multicount_gen:
   61.86 -  assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
   61.87 -  shows "setsum (\<lambda>i. (card {j\<in>t. R i j})) s = setsum k t" (is "?l = ?r")
   61.88 -proof-
   61.89 -  have "?l = setsum (\<lambda>i. setsum (\<lambda>x.1) {j\<in>t. R i j}) s" by auto
   61.90 -  also have "\<dots> = ?r" unfolding setsum_setsum_restrict[OF assms(1-2)]
   61.91 -    using assms(3) by auto
   61.92 -  finally show ?thesis .
   61.93 -qed
   61.94 -
   61.95 -lemma setsum_multicount:
   61.96 -  assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
   61.97 -  shows "setsum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
   61.98 -proof-
   61.99 -  have "?l = setsum (\<lambda>i. k) T" by(rule setsum_multicount_gen)(auto simp:assms)
  61.100 -  also have "\<dots> = ?r" by(simp add: mult_commute)
  61.101 -  finally show ?thesis by auto
  61.102 -qed
  61.103 -
  61.104  lemma setsum_nat_group: "(\<Sum>m<n::nat. setsum f {m * k ..< m*k + k}) = setsum f {..< n * k}"
  61.105    apply (subgoal_tac "k = 0 | 0 < k", auto)
  61.106    apply (induct "n")
  61.107 @@ -1587,11 +1523,11 @@
  61.108  
  61.109  lemma nested_setsum_swap:
  61.110       "(\<Sum>i = 0..n. (\<Sum>j = 0..<i. a i j)) = (\<Sum>j = 0..<n. \<Sum>i = Suc j..n. a i j)"
  61.111 -  by (induction n) (auto simp: setsum_addf)
  61.112 +  by (induction n) (auto simp: setsum.distrib)
  61.113  
  61.114  lemma nested_setsum_swap':
  61.115       "(\<Sum>i\<le>n. (\<Sum>j<i. a i j)) = (\<Sum>j<n. \<Sum>i = Suc j..n. a i j)"
  61.116 -  by (induction n) (auto simp: setsum_addf)
  61.117 +  by (induction n) (auto simp: setsum.distrib)
  61.118  
  61.119  lemma setsum_zero_power [simp]:
  61.120    fixes c :: "nat \<Rightarrow> 'a::division_ring"
  61.121 @@ -1642,7 +1578,7 @@
  61.122    have
  61.123      "(\<Sum>i\<in>{..<n}. a+?I i*d) =
  61.124       ((\<Sum>i\<in>{..<n}. a) + (\<Sum>i\<in>{..<n}. ?I i*d))"
  61.125 -    by (rule setsum_addf)
  61.126 +    by (rule setsum.distrib)
  61.127    also from ngt1 have "\<dots> = ?n*a + (\<Sum>i\<in>{..<n}. ?I i*d)" by simp
  61.128    also from ngt1 have "\<dots> = (?n*a + d*(\<Sum>i\<in>{1..<n}. ?I i))"
  61.129      unfolding One_nat_def
  61.130 @@ -1791,15 +1727,4 @@
  61.131      by auto
  61.132  qed
  61.133  
  61.134 -lemma setprod_power_distrib:
  61.135 -  fixes f :: "'a \<Rightarrow> 'b::comm_semiring_1"
  61.136 -  shows "setprod f A ^ n = setprod (\<lambda>x. (f x)^n) A"
  61.137 -proof (cases "finite A") 
  61.138 -  case True then show ?thesis 
  61.139 -    by (induct A rule: finite_induct) (auto simp add: power_mult_distrib)
  61.140 -next
  61.141 -  case False then show ?thesis 
  61.142 -    by (metis setprod_infinite power_one)
  61.143 -qed
  61.144 -
  61.145  end
    62.1 --- a/src/HOL/Transcendental.thy	Fri Jun 27 22:08:55 2014 +0200
    62.2 +++ b/src/HOL/Transcendental.thy	Sat Jun 28 09:16:42 2014 +0200
    62.3 @@ -237,7 +237,7 @@
    62.4      have "?s sums y" using sums_if'[OF `f sums y`] .
    62.5      from this[unfolded sums_def, THEN LIMSEQ_Suc]
    62.6      have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
    62.7 -      by (simp add: lessThan_Suc_eq_insert_0 image_iff setsum_reindex if_eq sums_def cong del: if_cong)
    62.8 +      by (simp add: lessThan_Suc_eq_insert_0 image_iff setsum.reindex if_eq sums_def cong del: if_cong)
    62.9    }
   62.10    from sums_add[OF g_sums this] show ?thesis unfolding if_sum .
   62.11  qed
   62.12 @@ -478,7 +478,7 @@
   62.13    apply (subst sumr_diff_mult_const2)
   62.14    apply simp
   62.15    apply (simp only: lemma_termdiff1 setsum_right_distrib)
   62.16 -  apply (rule setsum_cong [OF refl])
   62.17 +  apply (rule setsum.cong [OF refl])
   62.18    apply (simp add: less_iff_Suc_add)
   62.19    apply (clarify)
   62.20    apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
   62.21 @@ -1071,7 +1071,7 @@
   62.22    also have "(\<Sum>i\<le>Suc n. real i *\<^sub>R (S x i * S y (Suc n-i))) +
   62.23               (\<Sum>i\<le>Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) =
   62.24               (\<Sum>i\<le>Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n-i)))"
   62.25 -    by (simp only: setsum_addf [symmetric] scaleR_left_distrib [symmetric]
   62.26 +    by (simp only: setsum.distrib [symmetric] scaleR_left_distrib [symmetric]
   62.27                     real_of_nat_add [symmetric]) simp
   62.28    also have "\<dots> = real (Suc n) *\<^sub>R (\<Sum>i\<le>Suc n. S x i * S y (Suc n-i))"
   62.29      by (simp only: scaleR_right.setsum)
    63.1 --- a/src/HOL/UNITY/Comp/AllocBase.thy	Fri Jun 27 22:08:55 2014 +0200
    63.2 +++ b/src/HOL/UNITY/Comp/AllocBase.thy	Sat Jun 28 09:16:42 2014 +0200
    63.3 @@ -51,12 +51,12 @@
    63.4  
    63.5  (** setsum **)
    63.6  
    63.7 -declare setsum_cong [cong]
    63.8 +declare setsum.cong [cong]
    63.9  
   63.10  lemma bag_of_sublist_lemma:
   63.11       "(\<Sum>i\<in> A Int lessThan k. {#if i<k then f i else g i#}) =  
   63.12        (\<Sum>i\<in> A Int lessThan k. {#f i#})"
   63.13 -by (rule setsum_cong, auto)
   63.14 +by (rule setsum.cong, auto)
   63.15  
   63.16  lemma bag_of_sublist:
   63.17       "bag_of (sublist l A) =  
   63.18 @@ -72,7 +72,7 @@
   63.19        bag_of (sublist l A) + bag_of (sublist l B)"
   63.20  apply (subgoal_tac "A Int B Int {..<length l} =
   63.21                      (A Int {..<length l}) Int (B Int {..<length l}) ")
   63.22 -apply (simp add: bag_of_sublist Int_Un_distrib2 setsum_Un_Int, blast)
   63.23 +apply (simp add: bag_of_sublist Int_Un_distrib2 setsum.union_inter, blast)
   63.24  done
   63.25  
   63.26  lemma bag_of_sublist_Un_disjoint:
   63.27 @@ -87,7 +87,7 @@
   63.28            (\<Sum>i\<in>I. bag_of (sublist l (A i)))"
   63.29  apply (simp del: UN_simps 
   63.30              add: UN_simps [symmetric] add: bag_of_sublist)
   63.31 -apply (subst setsum_UN_disjoint, auto)
   63.32 +apply (subst setsum.UNION_disjoint, auto)
   63.33  done
   63.34  
   63.35  end
    64.1 --- a/src/HOL/ex/HarmonicSeries.thy	Fri Jun 27 22:08:55 2014 +0200
    64.2 +++ b/src/HOL/ex/HarmonicSeries.thy	Sat Jun 28 09:16:42 2014 +0200
    64.3 @@ -196,7 +196,7 @@
    64.4            "finite {1..(2::nat)^M}" and "finite {(2::nat)^M+1..(2::nat)^(Suc M)}"
    64.5            by auto
    64.6          ultimately show ?thesis
    64.7 -          by (auto intro: setsum_Un_disjoint)
    64.8 +          by (auto intro: setsum.union_disjoint)
    64.9        qed
   64.10        moreover
   64.11        {
    65.1 --- a/src/HOL/ex/Transfer_Int_Nat.thy	Fri Jun 27 22:08:55 2014 +0200
    65.2 +++ b/src/HOL/ex/Transfer_Int_Nat.thy	Sat Jun 28 09:16:42 2014 +0200
    65.3 @@ -110,7 +110,8 @@
    65.4    apply (intro rel_funI)
    65.5    apply (erule (1) bi_unique_rel_set_lemma)
    65.6    apply (simp add: setsum.reindex int_setsum ZN_def rel_fun_def)
    65.7 -  apply (rule setsum_cong2, simp)
    65.8 +  apply (rule setsum.cong)
    65.9 +  apply simp_all
   65.10    done
   65.11  
   65.12  text {* For derived operations, we can use the @{text "transfer_prover"}