1.1 --- /dev/null Thu Jan 01 00:00:00 1970 +0000
1.2 +++ b/src/HOL/Groups_Big.thy Sun Dec 15 15:10:14 2013 +0100
1.3 @@ -0,0 +1,1379 @@
1.4 +(* Title: HOL/Groups_Big.thy
1.5 + Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
1.6 + with contributions by Jeremy Avigad
1.7 +*)
1.8 +
1.9 +header {* Big sum and product over finite (non-empty) sets *}
1.10 +
1.11 +theory Groups_Big
1.12 +imports Finite_Set
1.13 +begin
1.14 +
1.15 +subsection {* Generic monoid operation over a set *}
1.16 +
1.17 +no_notation times (infixl "*" 70)
1.18 +no_notation Groups.one ("1")
1.19 +
1.20 +locale comm_monoid_set = comm_monoid
1.21 +begin
1.22 +
1.23 +interpretation comp_fun_commute f
1.24 + by default (simp add: fun_eq_iff left_commute)
1.25 +
1.26 +interpretation comp_fun_commute "f \<circ> g"
1.27 + by (rule comp_comp_fun_commute)
1.28 +
1.29 +definition F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
1.30 +where
1.31 + eq_fold: "F g A = Finite_Set.fold (f \<circ> g) 1 A"
1.32 +
1.33 +lemma infinite [simp]:
1.34 + "\<not> finite A \<Longrightarrow> F g A = 1"
1.35 + by (simp add: eq_fold)
1.36 +
1.37 +lemma empty [simp]:
1.38 + "F g {} = 1"
1.39 + by (simp add: eq_fold)
1.40 +
1.41 +lemma insert [simp]:
1.42 + assumes "finite A" and "x \<notin> A"
1.43 + shows "F g (insert x A) = g x * F g A"
1.44 + using assms by (simp add: eq_fold)
1.45 +
1.46 +lemma remove:
1.47 + assumes "finite A" and "x \<in> A"
1.48 + shows "F g A = g x * F g (A - {x})"
1.49 +proof -
1.50 + from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B"
1.51 + by (auto dest: mk_disjoint_insert)
1.52 + moreover from `finite A` A have "finite B" by simp
1.53 + ultimately show ?thesis by simp
1.54 +qed
1.55 +
1.56 +lemma insert_remove:
1.57 + assumes "finite A"
1.58 + shows "F g (insert x A) = g x * F g (A - {x})"
1.59 + using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb)
1.60 +
1.61 +lemma neutral:
1.62 + assumes "\<forall>x\<in>A. g x = 1"
1.63 + shows "F g A = 1"
1.64 + using assms by (induct A rule: infinite_finite_induct) simp_all
1.65 +
1.66 +lemma neutral_const [simp]:
1.67 + "F (\<lambda>_. 1) A = 1"
1.68 + by (simp add: neutral)
1.69 +
1.70 +lemma union_inter:
1.71 + assumes "finite A" and "finite B"
1.72 + shows "F g (A \<union> B) * F g (A \<inter> B) = F g A * F g B"
1.73 + -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
1.74 +using assms proof (induct A)
1.75 + case empty then show ?case by simp
1.76 +next
1.77 + case (insert x A) then show ?case
1.78 + by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute)
1.79 +qed
1.80 +
1.81 +corollary union_inter_neutral:
1.82 + assumes "finite A" and "finite B"
1.83 + and I0: "\<forall>x \<in> A \<inter> B. g x = 1"
1.84 + shows "F g (A \<union> B) = F g A * F g B"
1.85 + using assms by (simp add: union_inter [symmetric] neutral)
1.86 +
1.87 +corollary union_disjoint:
1.88 + assumes "finite A" and "finite B"
1.89 + assumes "A \<inter> B = {}"
1.90 + shows "F g (A \<union> B) = F g A * F g B"
1.91 + using assms by (simp add: union_inter_neutral)
1.92 +
1.93 +lemma subset_diff:
1.94 + assumes "B \<subseteq> A" and "finite A"
1.95 + shows "F g A = F g (A - B) * F g B"
1.96 +proof -
1.97 + from assms have "finite (A - B)" by auto
1.98 + moreover from assms have "finite B" by (rule finite_subset)
1.99 + moreover from assms have "(A - B) \<inter> B = {}" by auto
1.100 + ultimately have "F g (A - B \<union> B) = F g (A - B) * F g B" by (rule union_disjoint)
1.101 + moreover from assms have "A \<union> B = A" by auto
1.102 + ultimately show ?thesis by simp
1.103 +qed
1.104 +
1.105 +lemma reindex:
1.106 + assumes "inj_on h A"
1.107 + shows "F g (h ` A) = F (g \<circ> h) A"
1.108 +proof (cases "finite A")
1.109 + case True
1.110 + with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc)
1.111 +next
1.112 + case False with assms have "\<not> finite (h ` A)" by (blast dest: finite_imageD)
1.113 + with False show ?thesis by simp
1.114 +qed
1.115 +
1.116 +lemma cong:
1.117 + assumes "A = B"
1.118 + assumes g_h: "\<And>x. x \<in> B \<Longrightarrow> g x = h x"
1.119 + shows "F g A = F h B"
1.120 +proof (cases "finite A")
1.121 + case True
1.122 + then have "\<And>C. C \<subseteq> A \<longrightarrow> (\<forall>x\<in>C. g x = h x) \<longrightarrow> F g C = F h C"
1.123 + proof induct
1.124 + case empty then show ?case by simp
1.125 + next
1.126 + case (insert x F) then show ?case apply -
1.127 + apply (simp add: subset_insert_iff, clarify)
1.128 + apply (subgoal_tac "finite C")
1.129 + prefer 2 apply (blast dest: finite_subset [rotated])
1.130 + apply (subgoal_tac "C = insert x (C - {x})")
1.131 + prefer 2 apply blast
1.132 + apply (erule ssubst)
1.133 + apply (simp add: Ball_def del: insert_Diff_single)
1.134 + done
1.135 + qed
1.136 + with `A = B` g_h show ?thesis by simp
1.137 +next
1.138 + case False
1.139 + with `A = B` show ?thesis by simp
1.140 +qed
1.141 +
1.142 +lemma strong_cong [cong]:
1.143 + assumes "A = B" "\<And>x. x \<in> B =simp=> g x = h x"
1.144 + shows "F (\<lambda>x. g x) A = F (\<lambda>x. h x) B"
1.145 + by (rule cong) (insert assms, simp_all add: simp_implies_def)
1.146 +
1.147 +lemma UNION_disjoint:
1.148 + assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
1.149 + and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
1.150 + shows "F g (UNION I A) = F (\<lambda>x. F g (A x)) I"
1.151 +apply (insert assms)
1.152 +apply (induct rule: finite_induct)
1.153 +apply simp
1.154 +apply atomize
1.155 +apply (subgoal_tac "\<forall>i\<in>Fa. x \<noteq> i")
1.156 + prefer 2 apply blast
1.157 +apply (subgoal_tac "A x Int UNION Fa A = {}")
1.158 + prefer 2 apply blast
1.159 +apply (simp add: union_disjoint)
1.160 +done
1.161 +
1.162 +lemma Union_disjoint:
1.163 + assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {}"
1.164 + shows "F g (Union C) = F (F g) C"
1.165 +proof cases
1.166 + assume "finite C"
1.167 + from UNION_disjoint [OF this assms]
1.168 + show ?thesis
1.169 + by (simp add: SUP_def)
1.170 +qed (auto dest: finite_UnionD intro: infinite)
1.171 +
1.172 +lemma distrib:
1.173 + "F (\<lambda>x. g x * h x) A = F g A * F h A"
1.174 + using assms by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute)
1.175 +
1.176 +lemma Sigma:
1.177 + "finite A \<Longrightarrow> \<forall>x\<in>A. finite (B x) \<Longrightarrow> F (\<lambda>x. F (g x) (B x)) A = F (split g) (SIGMA x:A. B x)"
1.178 +apply (subst Sigma_def)
1.179 +apply (subst UNION_disjoint, assumption, simp)
1.180 + apply blast
1.181 +apply (rule cong)
1.182 +apply rule
1.183 +apply (simp add: fun_eq_iff)
1.184 +apply (subst UNION_disjoint, simp, simp)
1.185 + apply blast
1.186 +apply (simp add: comp_def)
1.187 +done
1.188 +
1.189 +lemma related:
1.190 + assumes Re: "R 1 1"
1.191 + and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 * y1) (x2 * y2)"
1.192 + and fS: "finite S" and Rfg: "\<forall>x\<in>S. R (h x) (g x)"
1.193 + shows "R (F h S) (F g S)"
1.194 + using fS by (rule finite_subset_induct) (insert assms, auto)
1.195 +
1.196 +lemma eq_general:
1.197 + assumes h: "\<forall>y\<in>S'. \<exists>!x. x \<in> S \<and> h x = y"
1.198 + and f12: "\<forall>x\<in>S. h x \<in> S' \<and> f2 (h x) = f1 x"
1.199 + shows "F f1 S = F f2 S'"
1.200 +proof-
1.201 + from h f12 have hS: "h ` S = S'" by blast
1.202 + {fix x y assume H: "x \<in> S" "y \<in> S" "h x = h y"
1.203 + from f12 h H have "x = y" by auto }
1.204 + hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast
1.205 + from f12 have th: "\<And>x. x \<in> S \<Longrightarrow> (f2 \<circ> h) x = f1 x" by auto
1.206 + from hS have "F f2 S' = F f2 (h ` S)" by simp
1.207 + also have "\<dots> = F (f2 o h) S" using reindex [OF hinj, of f2] .
1.208 + also have "\<dots> = F f1 S " using th cong [of _ _ "f2 o h" f1]
1.209 + by blast
1.210 + finally show ?thesis ..
1.211 +qed
1.212 +
1.213 +lemma eq_general_reverses:
1.214 + assumes kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
1.215 + and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = j x"
1.216 + shows "F j S = F g T"
1.217 + (* metis solves it, but not yet available here *)
1.218 + apply (rule eq_general [of T S h g j])
1.219 + apply (rule ballI)
1.220 + apply (frule kh)
1.221 + apply (rule ex1I[])
1.222 + apply blast
1.223 + apply clarsimp
1.224 + apply (drule hk) apply simp
1.225 + apply (rule sym)
1.226 + apply (erule conjunct1[OF conjunct2[OF hk]])
1.227 + apply (rule ballI)
1.228 + apply (drule hk)
1.229 + apply blast
1.230 + done
1.231 +
1.232 +lemma mono_neutral_cong_left:
1.233 + assumes "finite T" and "S \<subseteq> T" and "\<forall>i \<in> T - S. h i = 1"
1.234 + and "\<And>x. x \<in> S \<Longrightarrow> g x = h x" shows "F g S = F h T"
1.235 +proof-
1.236 + have eq: "T = S \<union> (T - S)" using `S \<subseteq> T` by blast
1.237 + have d: "S \<inter> (T - S) = {}" using `S \<subseteq> T` by blast
1.238 + from `finite T` `S \<subseteq> T` have f: "finite S" "finite (T - S)"
1.239 + by (auto intro: finite_subset)
1.240 + show ?thesis using assms(4)
1.241 + by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)])
1.242 +qed
1.243 +
1.244 +lemma mono_neutral_cong_right:
1.245 + "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1; \<And>x. x \<in> S \<Longrightarrow> g x = h x \<rbrakk>
1.246 + \<Longrightarrow> F g T = F h S"
1.247 + by (auto intro!: mono_neutral_cong_left [symmetric])
1.248 +
1.249 +lemma mono_neutral_left:
1.250 + "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g S = F g T"
1.251 + by (blast intro: mono_neutral_cong_left)
1.252 +
1.253 +lemma mono_neutral_right:
1.254 + "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g T = F g S"
1.255 + by (blast intro!: mono_neutral_left [symmetric])
1.256 +
1.257 +lemma delta:
1.258 + assumes fS: "finite S"
1.259 + shows "F (\<lambda>k. if k = a then b k else 1) S = (if a \<in> S then b a else 1)"
1.260 +proof-
1.261 + let ?f = "(\<lambda>k. if k=a then b k else 1)"
1.262 + { assume a: "a \<notin> S"
1.263 + hence "\<forall>k\<in>S. ?f k = 1" by simp
1.264 + hence ?thesis using a by simp }
1.265 + moreover
1.266 + { assume a: "a \<in> S"
1.267 + let ?A = "S - {a}"
1.268 + let ?B = "{a}"
1.269 + have eq: "S = ?A \<union> ?B" using a by blast
1.270 + have dj: "?A \<inter> ?B = {}" by simp
1.271 + from fS have fAB: "finite ?A" "finite ?B" by auto
1.272 + have "F ?f S = F ?f ?A * F ?f ?B"
1.273 + using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]]
1.274 + by simp
1.275 + then have ?thesis using a by simp }
1.276 + ultimately show ?thesis by blast
1.277 +qed
1.278 +
1.279 +lemma delta':
1.280 + assumes fS: "finite S"
1.281 + shows "F (\<lambda>k. if a = k then b k else 1) S = (if a \<in> S then b a else 1)"
1.282 + using delta [OF fS, of a b, symmetric] by (auto intro: cong)
1.283 +
1.284 +lemma If_cases:
1.285 + fixes P :: "'b \<Rightarrow> bool" and g h :: "'b \<Rightarrow> 'a"
1.286 + assumes fA: "finite A"
1.287 + shows "F (\<lambda>x. if P x then h x else g x) A =
1.288 + F h (A \<inter> {x. P x}) * F g (A \<inter> - {x. P x})"
1.289 +proof -
1.290 + have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}"
1.291 + "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
1.292 + by blast+
1.293 + from fA
1.294 + have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
1.295 + let ?g = "\<lambda>x. if P x then h x else g x"
1.296 + from union_disjoint [OF f a(2), of ?g] a(1)
1.297 + show ?thesis
1.298 + by (subst (1 2) cong) simp_all
1.299 +qed
1.300 +
1.301 +lemma cartesian_product:
1.302 + "F (\<lambda>x. F (g x) B) A = F (split g) (A <*> B)"
1.303 +apply (rule sym)
1.304 +apply (cases "finite A")
1.305 + apply (cases "finite B")
1.306 + apply (simp add: Sigma)
1.307 + apply (cases "A={}", simp)
1.308 + apply simp
1.309 +apply (auto intro: infinite dest: finite_cartesian_productD2)
1.310 +apply (cases "B = {}") apply (auto intro: infinite dest: finite_cartesian_productD1)
1.311 +done
1.312 +
1.313 +end
1.314 +
1.315 +notation times (infixl "*" 70)
1.316 +notation Groups.one ("1")
1.317 +
1.318 +
1.319 +subsection {* Generalized summation over a set *}
1.320 +
1.321 +context comm_monoid_add
1.322 +begin
1.323 +
1.324 +definition setsum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
1.325 +where
1.326 + "setsum = comm_monoid_set.F plus 0"
1.327 +
1.328 +sublocale setsum!: comm_monoid_set plus 0
1.329 +where
1.330 + "comm_monoid_set.F plus 0 = setsum"
1.331 +proof -
1.332 + show "comm_monoid_set plus 0" ..
1.333 + then interpret setsum!: comm_monoid_set plus 0 .
1.334 + from setsum_def show "comm_monoid_set.F plus 0 = setsum" by rule
1.335 +qed
1.336 +
1.337 +abbreviation
1.338 + Setsum ("\<Sum>_" [1000] 999) where
1.339 + "\<Sum>A \<equiv> setsum (%x. x) A"
1.340 +
1.341 +end
1.342 +
1.343 +text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
1.344 +written @{text"\<Sum>x\<in>A. e"}. *}
1.345 +
1.346 +syntax
1.347 + "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3SUM _:_. _)" [0, 51, 10] 10)
1.348 +syntax (xsymbols)
1.349 + "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
1.350 +syntax (HTML output)
1.351 + "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add" ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
1.352 +
1.353 +translations -- {* Beware of argument permutation! *}
1.354 + "SUM i:A. b" == "CONST setsum (%i. b) A"
1.355 + "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A"
1.356 +
1.357 +text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
1.358 + @{text"\<Sum>x|P. e"}. *}
1.359 +
1.360 +syntax
1.361 + "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
1.362 +syntax (xsymbols)
1.363 + "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
1.364 +syntax (HTML output)
1.365 + "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
1.366 +
1.367 +translations
1.368 + "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
1.369 + "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
1.370 +
1.371 +print_translation {*
1.372 +let
1.373 + fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
1.374 + if x <> y then raise Match
1.375 + else
1.376 + let
1.377 + val x' = Syntax_Trans.mark_bound_body (x, Tx);
1.378 + val t' = subst_bound (x', t);
1.379 + val P' = subst_bound (x', P);
1.380 + in
1.381 + Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
1.382 + end
1.383 + | setsum_tr' _ = raise Match;
1.384 +in [(@{const_syntax setsum}, K setsum_tr')] end
1.385 +*}
1.386 +
1.387 +text {* TODO These are candidates for generalization *}
1.388 +
1.389 +context comm_monoid_add
1.390 +begin
1.391 +
1.392 +lemma setsum_reindex_id:
1.393 + "inj_on f B ==> setsum f B = setsum id (f ` B)"
1.394 + by (simp add: setsum.reindex)
1.395 +
1.396 +lemma setsum_reindex_nonzero:
1.397 + assumes fS: "finite S"
1.398 + 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"
1.399 + shows "setsum h (f ` S) = setsum (h \<circ> f) S"
1.400 +using nz proof (induct rule: finite_induct [OF fS])
1.401 + case 1 thus ?case by simp
1.402 +next
1.403 + case (2 x F)
1.404 + { assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
1.405 + then obtain y where y: "y \<in> F" "f x = f y" by auto
1.406 + from "2.hyps" y have xy: "x \<noteq> y" by auto
1.407 + from "2.prems" [of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp
1.408 + have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto
1.409 + also have "\<dots> = setsum (h o f) (insert x F)"
1.410 + unfolding setsum.insert[OF `finite F` `x\<notin>F`]
1.411 + using h0
1.412 + apply (simp cong del: setsum.strong_cong)
1.413 + apply (rule "2.hyps"(3))
1.414 + apply (rule_tac y="y" in "2.prems")
1.415 + apply simp_all
1.416 + done
1.417 + finally have ?case . }
1.418 + moreover
1.419 + { assume fxF: "f x \<notin> f ` F"
1.420 + have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)"
1.421 + using fxF "2.hyps" by simp
1.422 + also have "\<dots> = setsum (h o f) (insert x F)"
1.423 + unfolding setsum.insert[OF `finite F` `x\<notin>F`]
1.424 + apply (simp cong del: setsum.strong_cong)
1.425 + apply (rule cong [OF refl [of "op + (h (f x))"]])
1.426 + apply (rule "2.hyps"(3))
1.427 + apply (rule_tac y="y" in "2.prems")
1.428 + apply simp_all
1.429 + done
1.430 + finally have ?case . }
1.431 + ultimately show ?case by blast
1.432 +qed
1.433 +
1.434 +lemma setsum_cong2:
1.435 + "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A"
1.436 + by (auto intro: setsum.cong)
1.437 +
1.438 +lemma setsum_reindex_cong:
1.439 + "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|]
1.440 + ==> setsum h B = setsum g A"
1.441 + by (simp add: setsum.reindex)
1.442 +
1.443 +lemma setsum_restrict_set:
1.444 + assumes fA: "finite A"
1.445 + shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A"
1.446 +proof-
1.447 + from fA have fab: "finite (A \<inter> B)" by auto
1.448 + have aba: "A \<inter> B \<subseteq> A" by blast
1.449 + let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0"
1.450 + from setsum.mono_neutral_left [OF fA aba, of ?g]
1.451 + show ?thesis by simp
1.452 +qed
1.453 +
1.454 +lemma setsum_Union_disjoint:
1.455 + assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}"
1.456 + shows "setsum f (Union C) = setsum (setsum f) C"
1.457 + using assms by (fact setsum.Union_disjoint)
1.458 +
1.459 +lemma setsum_cartesian_product:
1.460 + "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)"
1.461 + by (fact setsum.cartesian_product)
1.462 +
1.463 +lemma setsum_UNION_zero:
1.464 + assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T"
1.465 + 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"
1.466 + shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S"
1.467 + using fSS f0
1.468 +proof(induct rule: finite_induct[OF fS])
1.469 + case 1 thus ?case by simp
1.470 +next
1.471 + case (2 T F)
1.472 + then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F"
1.473 + and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto
1.474 + from fTF have fUF: "finite (\<Union>F)" by auto
1.475 + from "2.prems" TF fTF
1.476 + show ?case
1.477 + by (auto simp add: H [symmetric] intro: setsum.union_inter_neutral [OF fTF(1) fUF, of f])
1.478 +qed
1.479 +
1.480 +text {* Commuting outer and inner summation *}
1.481 +
1.482 +lemma setsum_commute:
1.483 + "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)"
1.484 +proof (simp add: setsum_cartesian_product)
1.485 + have "(\<Sum>(x,y) \<in> A <*> B. f x y) =
1.486 + (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)"
1.487 + (is "?s = _")
1.488 + apply (simp add: setsum.reindex [where h = "%(i, j). (j, i)"] swap_inj_on)
1.489 + apply (simp add: split_def)
1.490 + done
1.491 + also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)"
1.492 + (is "_ = ?t")
1.493 + apply (simp add: swap_product)
1.494 + done
1.495 + finally show "?s = ?t" .
1.496 +qed
1.497 +
1.498 +lemma setsum_Plus:
1.499 + fixes A :: "'a set" and B :: "'b set"
1.500 + assumes fin: "finite A" "finite B"
1.501 + shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B"
1.502 +proof -
1.503 + have "A <+> B = Inl ` A \<union> Inr ` B" by auto
1.504 + moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
1.505 + by auto
1.506 + moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
1.507 + moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI)
1.508 + ultimately show ?thesis using fin by(simp add: setsum.union_disjoint setsum.reindex)
1.509 +qed
1.510 +
1.511 +end
1.512 +
1.513 +text {* TODO These are legacy *}
1.514 +
1.515 +lemma setsum_empty:
1.516 + "setsum f {} = 0"
1.517 + by (fact setsum.empty)
1.518 +
1.519 +lemma setsum_insert:
1.520 + "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F"
1.521 + by (fact setsum.insert)
1.522 +
1.523 +lemma setsum_infinite:
1.524 + "~ finite A ==> setsum f A = 0"
1.525 + by (fact setsum.infinite)
1.526 +
1.527 +lemma setsum_reindex:
1.528 + "inj_on f B \<Longrightarrow> setsum h (f ` B) = setsum (h \<circ> f) B"
1.529 + by (fact setsum.reindex)
1.530 +
1.531 +lemma setsum_cong:
1.532 + "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B"
1.533 + by (fact setsum.cong)
1.534 +
1.535 +lemma strong_setsum_cong:
1.536 + "A = B ==> (!!x. x:B =simp=> f x = g x)
1.537 + ==> setsum (%x. f x) A = setsum (%x. g x) B"
1.538 + by (fact setsum.strong_cong)
1.539 +
1.540 +lemmas setsum_0 = setsum.neutral_const
1.541 +lemmas setsum_0' = setsum.neutral
1.542 +
1.543 +lemma setsum_Un_Int: "finite A ==> finite B ==>
1.544 + setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B"
1.545 + -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
1.546 + by (fact setsum.union_inter)
1.547 +
1.548 +lemma setsum_Un_disjoint: "finite A ==> finite B
1.549 + ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
1.550 + by (fact setsum.union_disjoint)
1.551 +
1.552 +lemma setsum_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>
1.553 + setsum f A = setsum f (A - B) + setsum f B"
1.554 + by (fact setsum.subset_diff)
1.555 +
1.556 +lemma setsum_mono_zero_left:
1.557 + "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 0 \<rbrakk> \<Longrightarrow> setsum f S = setsum f T"
1.558 + by (fact setsum.mono_neutral_left)
1.559 +
1.560 +lemmas setsum_mono_zero_right = setsum.mono_neutral_right
1.561 +
1.562 +lemma setsum_mono_zero_cong_left:
1.563 + "\<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>
1.564 + \<Longrightarrow> setsum f S = setsum g T"
1.565 + by (fact setsum.mono_neutral_cong_left)
1.566 +
1.567 +lemmas setsum_mono_zero_cong_right = setsum.mono_neutral_cong_right
1.568 +
1.569 +lemma setsum_delta: "finite S \<Longrightarrow>
1.570 + setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)"
1.571 + by (fact setsum.delta)
1.572 +
1.573 +lemma setsum_delta': "finite S \<Longrightarrow>
1.574 + setsum (\<lambda>k. if a = k then b k else 0) S = (if a\<in> S then b a else 0)"
1.575 + by (fact setsum.delta')
1.576 +
1.577 +lemma setsum_cases:
1.578 + assumes "finite A"
1.579 + shows "setsum (\<lambda>x. if P x then f x else g x) A =
1.580 + setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
1.581 + using assms by (fact setsum.If_cases)
1.582 +
1.583 +(*But we can't get rid of finite I. If infinite, although the rhs is 0,
1.584 + the lhs need not be, since UNION I A could still be finite.*)
1.585 +lemma setsum_UN_disjoint:
1.586 + assumes "finite I" and "ALL i:I. finite (A i)"
1.587 + and "ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}"
1.588 + shows "setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))"
1.589 + using assms by (fact setsum.UNION_disjoint)
1.590 +
1.591 +(*But we can't get rid of finite A. If infinite, although the lhs is 0,
1.592 + the rhs need not be, since SIGMA A B could still be finite.*)
1.593 +lemma setsum_Sigma:
1.594 + assumes "finite A" and "ALL x:A. finite (B x)"
1.595 + 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)"
1.596 + using assms by (fact setsum.Sigma)
1.597 +
1.598 +lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)"
1.599 + by (fact setsum.distrib)
1.600 +
1.601 +lemma setsum_Un_zero:
1.602 + "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 0 \<rbrakk> \<Longrightarrow>
1.603 + setsum f (S \<union> T) = setsum f S + setsum f T"
1.604 + by (fact setsum.union_inter_neutral)
1.605 +
1.606 +lemma setsum_eq_general_reverses:
1.607 + assumes fS: "finite S" and fT: "finite T"
1.608 + and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y"
1.609 + and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x"
1.610 + shows "setsum f S = setsum g T"
1.611 + using kh hk by (fact setsum.eq_general_reverses)
1.612 +
1.613 +
1.614 +subsubsection {* Properties in more restricted classes of structures *}
1.615 +
1.616 +lemma setsum_Un: "finite A ==> finite B ==>
1.617 + (setsum f (A Un B) :: 'a :: ab_group_add) =
1.618 + setsum f A + setsum f B - setsum f (A Int B)"
1.619 +by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
1.620 +
1.621 +lemma setsum_Un2:
1.622 + assumes "finite (A \<union> B)"
1.623 + shows "setsum f (A \<union> B) = setsum f (A - B) + setsum f (B - A) + setsum f (A \<inter> B)"
1.624 +proof -
1.625 + have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
1.626 + by auto
1.627 + with assms show ?thesis by simp (subst setsum_Un_disjoint, auto)+
1.628 +qed
1.629 +
1.630 +lemma setsum_diff1: "finite A \<Longrightarrow>
1.631 + (setsum f (A - {a}) :: ('a::ab_group_add)) =
1.632 + (if a:A then setsum f A - f a else setsum f A)"
1.633 +by (erule finite_induct) (auto simp add: insert_Diff_if)
1.634 +
1.635 +lemma setsum_diff:
1.636 + assumes le: "finite A" "B \<subseteq> A"
1.637 + shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
1.638 +proof -
1.639 + from le have finiteB: "finite B" using finite_subset by auto
1.640 + show ?thesis using finiteB le
1.641 + proof induct
1.642 + case empty
1.643 + thus ?case by auto
1.644 + next
1.645 + case (insert x F)
1.646 + thus ?case using le finiteB
1.647 + by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb)
1.648 + qed
1.649 +qed
1.650 +
1.651 +lemma setsum_mono:
1.652 + assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
1.653 + shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"
1.654 +proof (cases "finite K")
1.655 + case True
1.656 + thus ?thesis using le
1.657 + proof induct
1.658 + case empty
1.659 + thus ?case by simp
1.660 + next
1.661 + case insert
1.662 + thus ?case using add_mono by fastforce
1.663 + qed
1.664 +next
1.665 + case False then show ?thesis by simp
1.666 +qed
1.667 +
1.668 +lemma setsum_strict_mono:
1.669 + fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
1.670 + assumes "finite A" "A \<noteq> {}"
1.671 + and "!!x. x:A \<Longrightarrow> f x < g x"
1.672 + shows "setsum f A < setsum g A"
1.673 + using assms
1.674 +proof (induct rule: finite_ne_induct)
1.675 + case singleton thus ?case by simp
1.676 +next
1.677 + case insert thus ?case by (auto simp: add_strict_mono)
1.678 +qed
1.679 +
1.680 +lemma setsum_strict_mono_ex1:
1.681 +fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
1.682 +assumes "finite A" and "ALL x:A. f x \<le> g x" and "EX a:A. f a < g a"
1.683 +shows "setsum f A < setsum g A"
1.684 +proof-
1.685 + from assms(3) obtain a where a: "a:A" "f a < g a" by blast
1.686 + have "setsum f A = setsum f ((A-{a}) \<union> {a})"
1.687 + by(simp add:insert_absorb[OF `a:A`])
1.688 + also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
1.689 + using `finite A` by(subst setsum_Un_disjoint) auto
1.690 + also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
1.691 + by(rule setsum_mono)(simp add: assms(2))
1.692 + also have "setsum f {a} < setsum g {a}" using a by simp
1.693 + also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
1.694 + using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto
1.695 + also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`])
1.696 + finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono)
1.697 +qed
1.698 +
1.699 +lemma setsum_negf:
1.700 + "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A"
1.701 +proof (cases "finite A")
1.702 + case True thus ?thesis by (induct set: finite) auto
1.703 +next
1.704 + case False thus ?thesis by simp
1.705 +qed
1.706 +
1.707 +lemma setsum_subtractf:
1.708 + "setsum (%x. ((f x)::'a::ab_group_add) - g x) A =
1.709 + setsum f A - setsum g A"
1.710 + using setsum_addf [of f "- g" A] by (simp add: setsum_negf)
1.711 +
1.712 +lemma setsum_nonneg:
1.713 + assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
1.714 + shows "0 \<le> setsum f A"
1.715 +proof (cases "finite A")
1.716 + case True thus ?thesis using nn
1.717 + proof induct
1.718 + case empty then show ?case by simp
1.719 + next
1.720 + case (insert x F)
1.721 + then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono)
1.722 + with insert show ?case by simp
1.723 + qed
1.724 +next
1.725 + case False thus ?thesis by simp
1.726 +qed
1.727 +
1.728 +lemma setsum_nonpos:
1.729 + assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
1.730 + shows "setsum f A \<le> 0"
1.731 +proof (cases "finite A")
1.732 + case True thus ?thesis using np
1.733 + proof induct
1.734 + case empty then show ?case by simp
1.735 + next
1.736 + case (insert x F)
1.737 + then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono)
1.738 + with insert show ?case by simp
1.739 + qed
1.740 +next
1.741 + case False thus ?thesis by simp
1.742 +qed
1.743 +
1.744 +lemma setsum_nonneg_leq_bound:
1.745 + fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
1.746 + assumes "finite s" "\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0" "(\<Sum>i \<in> s. f i) = B" "i \<in> s"
1.747 + shows "f i \<le> B"
1.748 +proof -
1.749 + have "0 \<le> (\<Sum> i \<in> s - {i}. f i)" and "0 \<le> f i"
1.750 + using assms by (auto intro!: setsum_nonneg)
1.751 + moreover
1.752 + have "(\<Sum> i \<in> s - {i}. f i) + f i = B"
1.753 + using assms by (simp add: setsum_diff1)
1.754 + ultimately show ?thesis by auto
1.755 +qed
1.756 +
1.757 +lemma setsum_nonneg_0:
1.758 + fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
1.759 + assumes "finite s" and pos: "\<And> i. i \<in> s \<Longrightarrow> f i \<ge> 0"
1.760 + and "(\<Sum> i \<in> s. f i) = 0" and i: "i \<in> s"
1.761 + shows "f i = 0"
1.762 + using setsum_nonneg_leq_bound[OF assms] pos[OF i] by auto
1.763 +
1.764 +lemma setsum_mono2:
1.765 +fixes f :: "'a \<Rightarrow> 'b :: ordered_comm_monoid_add"
1.766 +assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"
1.767 +shows "setsum f A \<le> setsum f B"
1.768 +proof -
1.769 + have "setsum f A \<le> setsum f A + setsum f (B-A)"
1.770 + by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def)
1.771 + also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin]
1.772 + by (simp add:setsum_Un_disjoint del:Un_Diff_cancel)
1.773 + also have "A \<union> (B-A) = B" using sub by blast
1.774 + finally show ?thesis .
1.775 +qed
1.776 +
1.777 +lemma setsum_mono3: "finite B ==> A <= B ==>
1.778 + ALL x: B - A.
1.779 + 0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
1.780 + setsum f A <= setsum f B"
1.781 + apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)")
1.782 + apply (erule ssubst)
1.783 + apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)")
1.784 + apply simp
1.785 + apply (rule add_left_mono)
1.786 + apply (erule setsum_nonneg)
1.787 + apply (subst setsum_Un_disjoint [THEN sym])
1.788 + apply (erule finite_subset, assumption)
1.789 + apply (rule finite_subset)
1.790 + prefer 2
1.791 + apply assumption
1.792 + apply (auto simp add: sup_absorb2)
1.793 +done
1.794 +
1.795 +lemma setsum_right_distrib:
1.796 + fixes f :: "'a => ('b::semiring_0)"
1.797 + shows "r * setsum f A = setsum (%n. r * f n) A"
1.798 +proof (cases "finite A")
1.799 + case True
1.800 + thus ?thesis
1.801 + proof induct
1.802 + case empty thus ?case by simp
1.803 + next
1.804 + case (insert x A) thus ?case by (simp add: distrib_left)
1.805 + qed
1.806 +next
1.807 + case False thus ?thesis by simp
1.808 +qed
1.809 +
1.810 +lemma setsum_left_distrib:
1.811 + "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)"
1.812 +proof (cases "finite A")
1.813 + case True
1.814 + then show ?thesis
1.815 + proof induct
1.816 + case empty thus ?case by simp
1.817 + next
1.818 + case (insert x A) thus ?case by (simp add: distrib_right)
1.819 + qed
1.820 +next
1.821 + case False thus ?thesis by simp
1.822 +qed
1.823 +
1.824 +lemma setsum_divide_distrib:
1.825 + "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)"
1.826 +proof (cases "finite A")
1.827 + case True
1.828 + then show ?thesis
1.829 + proof induct
1.830 + case empty thus ?case by simp
1.831 + next
1.832 + case (insert x A) thus ?case by (simp add: add_divide_distrib)
1.833 + qed
1.834 +next
1.835 + case False thus ?thesis by simp
1.836 +qed
1.837 +
1.838 +lemma setsum_abs[iff]:
1.839 + fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
1.840 + shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A"
1.841 +proof (cases "finite A")
1.842 + case True
1.843 + thus ?thesis
1.844 + proof induct
1.845 + case empty thus ?case by simp
1.846 + next
1.847 + case (insert x A)
1.848 + thus ?case by (auto intro: abs_triangle_ineq order_trans)
1.849 + qed
1.850 +next
1.851 + case False thus ?thesis by simp
1.852 +qed
1.853 +
1.854 +lemma setsum_abs_ge_zero[iff]:
1.855 + fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
1.856 + shows "0 \<le> setsum (%i. abs(f i)) A"
1.857 +proof (cases "finite A")
1.858 + case True
1.859 + thus ?thesis
1.860 + proof induct
1.861 + case empty thus ?case by simp
1.862 + next
1.863 + case (insert x A) thus ?case by auto
1.864 + qed
1.865 +next
1.866 + case False thus ?thesis by simp
1.867 +qed
1.868 +
1.869 +lemma abs_setsum_abs[simp]:
1.870 + fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
1.871 + shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))"
1.872 +proof (cases "finite A")
1.873 + case True
1.874 + thus ?thesis
1.875 + proof induct
1.876 + case empty thus ?case by simp
1.877 + next
1.878 + case (insert a A)
1.879 + hence "\<bar>\<Sum>a\<in>insert a A. \<bar>f a\<bar>\<bar> = \<bar>\<bar>f a\<bar> + (\<Sum>a\<in>A. \<bar>f a\<bar>)\<bar>" by simp
1.880 + also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp
1.881 + also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>"
1.882 + by (simp del: abs_of_nonneg)
1.883 + also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp
1.884 + finally show ?case .
1.885 + qed
1.886 +next
1.887 + case False thus ?thesis by simp
1.888 +qed
1.889 +
1.890 +lemma setsum_diff1'[rule_format]:
1.891 + "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
1.892 +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))"])
1.893 +apply (auto simp add: insert_Diff_if add_ac)
1.894 +done
1.895 +
1.896 +lemma setsum_diff1_ring: assumes "finite A" "a \<in> A"
1.897 + shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
1.898 +unfolding setsum_diff1'[OF assms] by auto
1.899 +
1.900 +lemma setsum_product:
1.901 + fixes f :: "'a => ('b::semiring_0)"
1.902 + shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"
1.903 + by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute)
1.904 +
1.905 +lemma setsum_mult_setsum_if_inj:
1.906 +fixes f :: "'a => ('b::semiring_0)"
1.907 +shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==>
1.908 + setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
1.909 +by(auto simp: setsum_product setsum_cartesian_product
1.910 + intro!: setsum_reindex_cong[symmetric])
1.911 +
1.912 +lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a"
1.913 +apply (case_tac "finite A")
1.914 + prefer 2 apply simp
1.915 +apply (erule rev_mp)
1.916 +apply (erule finite_induct, auto)
1.917 +done
1.918 +
1.919 +lemma setsum_eq_0_iff [simp]:
1.920 + "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))"
1.921 + by (induct set: finite) auto
1.922 +
1.923 +lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow>
1.924 + setsum f A = Suc 0 \<longleftrightarrow> (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))"
1.925 +apply(erule finite_induct)
1.926 +apply (auto simp add:add_is_1)
1.927 +done
1.928 +
1.929 +lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]]
1.930 +
1.931 +lemma setsum_Un_nat: "finite A ==> finite B ==>
1.932 + (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)"
1.933 + -- {* For the natural numbers, we have subtraction. *}
1.934 +by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps)
1.935 +
1.936 +lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
1.937 + (if a:A then setsum f A - f a else setsum f A)"
1.938 +apply (case_tac "finite A")
1.939 + prefer 2 apply simp
1.940 +apply (erule finite_induct)
1.941 + apply (auto simp add: insert_Diff_if)
1.942 +apply (drule_tac a = a in mk_disjoint_insert, auto)
1.943 +done
1.944 +
1.945 +lemma setsum_diff_nat:
1.946 +assumes "finite B" and "B \<subseteq> A"
1.947 +shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)"
1.948 +using assms
1.949 +proof induct
1.950 + show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
1.951 +next
1.952 + fix F x assume finF: "finite F" and xnotinF: "x \<notin> F"
1.953 + and xFinA: "insert x F \<subseteq> A"
1.954 + and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F"
1.955 + from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp
1.956 + from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
1.957 + by (simp add: setsum_diff1_nat)
1.958 + from xFinA have "F \<subseteq> A" by simp
1.959 + with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp
1.960 + with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
1.961 + by simp
1.962 + from xnotinF have "A - insert x F = (A - F) - {x}" by auto
1.963 + with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x"
1.964 + by simp
1.965 + from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp
1.966 + with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)"
1.967 + by simp
1.968 + thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp
1.969 +qed
1.970 +
1.971 +lemma setsum_comp_morphism:
1.972 + assumes "h 0 = 0" and "\<And>x y. h (x + y) = h x + h y"
1.973 + shows "setsum (h \<circ> g) A = h (setsum g A)"
1.974 +proof (cases "finite A")
1.975 + case False then show ?thesis by (simp add: assms)
1.976 +next
1.977 + case True then show ?thesis by (induct A) (simp_all add: assms)
1.978 +qed
1.979 +
1.980 +
1.981 +subsubsection {* Cardinality as special case of @{const setsum} *}
1.982 +
1.983 +lemma card_eq_setsum:
1.984 + "card A = setsum (\<lambda>x. 1) A"
1.985 +proof -
1.986 + have "plus \<circ> (\<lambda>_. Suc 0) = (\<lambda>_. Suc)"
1.987 + by (simp add: fun_eq_iff)
1.988 + then have "Finite_Set.fold (plus \<circ> (\<lambda>_. Suc 0)) = Finite_Set.fold (\<lambda>_. Suc)"
1.989 + by (rule arg_cong)
1.990 + then have "Finite_Set.fold (plus \<circ> (\<lambda>_. Suc 0)) 0 A = Finite_Set.fold (\<lambda>_. Suc) 0 A"
1.991 + by (blast intro: fun_cong)
1.992 + then show ?thesis by (simp add: card.eq_fold setsum.eq_fold)
1.993 +qed
1.994 +
1.995 +lemma setsum_constant [simp]:
1.996 + "(\<Sum>x \<in> A. y) = of_nat (card A) * y"
1.997 +apply (cases "finite A")
1.998 +apply (erule finite_induct)
1.999 +apply (auto simp add: algebra_simps)
1.1000 +done
1.1001 +
1.1002 +lemma setsum_bounded:
1.1003 + assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
1.1004 + shows "setsum f A \<le> of_nat (card A) * K"
1.1005 +proof (cases "finite A")
1.1006 + case True
1.1007 + thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp
1.1008 +next
1.1009 + case False thus ?thesis by simp
1.1010 +qed
1.1011 +
1.1012 +lemma card_UN_disjoint:
1.1013 + assumes "finite I" and "\<forall>i\<in>I. finite (A i)"
1.1014 + and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
1.1015 + shows "card (UNION I A) = (\<Sum>i\<in>I. card(A i))"
1.1016 +proof -
1.1017 + have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)" by simp
1.1018 + with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant)
1.1019 +qed
1.1020 +
1.1021 +lemma card_Union_disjoint:
1.1022 + "finite C ==> (ALL A:C. finite A) ==>
1.1023 + (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
1.1024 + ==> card (Union C) = setsum card C"
1.1025 +apply (frule card_UN_disjoint [of C id])
1.1026 +apply (simp_all add: SUP_def id_def)
1.1027 +done
1.1028 +
1.1029 +
1.1030 +subsubsection {* Cardinality of products *}
1.1031 +
1.1032 +lemma card_SigmaI [simp]:
1.1033 + "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk>
1.1034 + \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"
1.1035 +by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant)
1.1036 +
1.1037 +(*
1.1038 +lemma SigmaI_insert: "y \<notin> A ==>
1.1039 + (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
1.1040 + by auto
1.1041 +*)
1.1042 +
1.1043 +lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)"
1.1044 + by (cases "finite A \<and> finite B")
1.1045 + (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2)
1.1046 +
1.1047 +lemma card_cartesian_product_singleton: "card({x} <*> A) = card(A)"
1.1048 +by (simp add: card_cartesian_product)
1.1049 +
1.1050 +
1.1051 +subsection {* Generalized product over a set *}
1.1052 +
1.1053 +context comm_monoid_mult
1.1054 +begin
1.1055 +
1.1056 +definition setprod :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
1.1057 +where
1.1058 + "setprod = comm_monoid_set.F times 1"
1.1059 +
1.1060 +sublocale setprod!: comm_monoid_set times 1
1.1061 +where
1.1062 + "comm_monoid_set.F times 1 = setprod"
1.1063 +proof -
1.1064 + show "comm_monoid_set times 1" ..
1.1065 + then interpret setprod!: comm_monoid_set times 1 .
1.1066 + from setprod_def show "comm_monoid_set.F times 1 = setprod" by rule
1.1067 +qed
1.1068 +
1.1069 +abbreviation
1.1070 + Setprod ("\<Prod>_" [1000] 999) where
1.1071 + "\<Prod>A \<equiv> setprod (\<lambda>x. x) A"
1.1072 +
1.1073 +end
1.1074 +
1.1075 +syntax
1.1076 + "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3PROD _:_. _)" [0, 51, 10] 10)
1.1077 +syntax (xsymbols)
1.1078 + "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
1.1079 +syntax (HTML output)
1.1080 + "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult" ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
1.1081 +
1.1082 +translations -- {* Beware of argument permutation! *}
1.1083 + "PROD i:A. b" == "CONST setprod (%i. b) A"
1.1084 + "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A"
1.1085 +
1.1086 +text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
1.1087 + @{text"\<Prod>x|P. e"}. *}
1.1088 +
1.1089 +syntax
1.1090 + "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
1.1091 +syntax (xsymbols)
1.1092 + "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
1.1093 +syntax (HTML output)
1.1094 + "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
1.1095 +
1.1096 +translations
1.1097 + "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
1.1098 + "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
1.1099 +
1.1100 +text {* TODO These are candidates for generalization *}
1.1101 +
1.1102 +context comm_monoid_mult
1.1103 +begin
1.1104 +
1.1105 +lemma setprod_reindex_id:
1.1106 + "inj_on f B ==> setprod f B = setprod id (f ` B)"
1.1107 + by (auto simp add: setprod.reindex)
1.1108 +
1.1109 +lemma setprod_reindex_cong:
1.1110 + "inj_on f A ==> B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A"
1.1111 + by (frule setprod.reindex, simp)
1.1112 +
1.1113 +lemma strong_setprod_reindex_cong:
1.1114 + assumes i: "inj_on f A"
1.1115 + and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x"
1.1116 + shows "setprod h B = setprod g A"
1.1117 +proof-
1.1118 + have "setprod h B = setprod (h o f) A"
1.1119 + by (simp add: B setprod.reindex [OF i, of h])
1.1120 + then show ?thesis apply simp
1.1121 + apply (rule setprod.cong)
1.1122 + apply simp
1.1123 + by (simp add: eq)
1.1124 +qed
1.1125 +
1.1126 +lemma setprod_Union_disjoint:
1.1127 + assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}"
1.1128 + shows "setprod f (Union C) = setprod (setprod f) C"
1.1129 + using assms by (fact setprod.Union_disjoint)
1.1130 +
1.1131 +text{*Here we can eliminate the finiteness assumptions, by cases.*}
1.1132 +lemma setprod_cartesian_product:
1.1133 + "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)"
1.1134 + by (fact setprod.cartesian_product)
1.1135 +
1.1136 +lemma setprod_Un2:
1.1137 + assumes "finite (A \<union> B)"
1.1138 + shows "setprod f (A \<union> B) = setprod f (A - B) * setprod f (B - A) * setprod f (A \<inter> B)"
1.1139 +proof -
1.1140 + have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"
1.1141 + by auto
1.1142 + with assms show ?thesis by simp (subst setprod.union_disjoint, auto)+
1.1143 +qed
1.1144 +
1.1145 +end
1.1146 +
1.1147 +text {* TODO These are legacy *}
1.1148 +
1.1149 +lemma setprod_empty: "setprod f {} = 1"
1.1150 + by (fact setprod.empty)
1.1151 +
1.1152 +lemma setprod_insert: "[| finite A; a \<notin> A |] ==>
1.1153 + setprod f (insert a A) = f a * setprod f A"
1.1154 + by (fact setprod.insert)
1.1155 +
1.1156 +lemma setprod_infinite: "~ finite A ==> setprod f A = 1"
1.1157 + by (fact setprod.infinite)
1.1158 +
1.1159 +lemma setprod_reindex:
1.1160 + "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B"
1.1161 + by (fact setprod.reindex)
1.1162 +
1.1163 +lemma setprod_cong:
1.1164 + "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B"
1.1165 + by (fact setprod.cong)
1.1166 +
1.1167 +lemma strong_setprod_cong:
1.1168 + "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B"
1.1169 + by (fact setprod.strong_cong)
1.1170 +
1.1171 +lemma setprod_Un_one:
1.1172 + "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 1 \<rbrakk>
1.1173 + \<Longrightarrow> setprod f (S \<union> T) = setprod f S * setprod f T"
1.1174 + by (fact setprod.union_inter_neutral)
1.1175 +
1.1176 +lemmas setprod_1 = setprod.neutral_const
1.1177 +lemmas setprod_1' = setprod.neutral
1.1178 +
1.1179 +lemma setprod_Un_Int: "finite A ==> finite B
1.1180 + ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B"
1.1181 + by (fact setprod.union_inter)
1.1182 +
1.1183 +lemma setprod_Un_disjoint: "finite A ==> finite B
1.1184 + ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
1.1185 + by (fact setprod.union_disjoint)
1.1186 +
1.1187 +lemma setprod_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow>
1.1188 + setprod f A = setprod f (A - B) * setprod f B"
1.1189 + by (fact setprod.subset_diff)
1.1190 +
1.1191 +lemma setprod_mono_one_left:
1.1192 + "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 1 \<rbrakk> \<Longrightarrow> setprod f S = setprod f T"
1.1193 + by (fact setprod.mono_neutral_left)
1.1194 +
1.1195 +lemmas setprod_mono_one_right = setprod.mono_neutral_right
1.1196 +
1.1197 +lemma setprod_mono_one_cong_left:
1.1198 + "\<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>
1.1199 + \<Longrightarrow> setprod f S = setprod g T"
1.1200 + by (fact setprod.mono_neutral_cong_left)
1.1201 +
1.1202 +lemmas setprod_mono_one_cong_right = setprod.mono_neutral_cong_right
1.1203 +
1.1204 +lemma setprod_delta: "finite S \<Longrightarrow>
1.1205 + setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)"
1.1206 + by (fact setprod.delta)
1.1207 +
1.1208 +lemma setprod_delta': "finite S \<Longrightarrow>
1.1209 + setprod (\<lambda>k. if a = k then b k else 1) S = (if a\<in> S then b a else 1)"
1.1210 + by (fact setprod.delta')
1.1211 +
1.1212 +lemma setprod_UN_disjoint:
1.1213 + "finite I ==> (ALL i:I. finite (A i)) ==>
1.1214 + (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
1.1215 + setprod f (UNION I A) = setprod (%i. setprod f (A i)) I"
1.1216 + by (fact setprod.UNION_disjoint)
1.1217 +
1.1218 +lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==>
1.1219 + (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) =
1.1220 + (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)"
1.1221 + by (fact setprod.Sigma)
1.1222 +
1.1223 +lemma setprod_timesf: "setprod (\<lambda>x. f x * g x) A = setprod f A * setprod g A"
1.1224 + by (fact setprod.distrib)
1.1225 +
1.1226 +
1.1227 +subsubsection {* Properties in more restricted classes of structures *}
1.1228 +
1.1229 +lemma setprod_zero:
1.1230 + "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0"
1.1231 +apply (induct set: finite, force, clarsimp)
1.1232 +apply (erule disjE, auto)
1.1233 +done
1.1234 +
1.1235 +lemma setprod_zero_iff[simp]: "finite A ==>
1.1236 + (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
1.1237 + (EX x: A. f x = 0)"
1.1238 +by (erule finite_induct, auto simp:no_zero_divisors)
1.1239 +
1.1240 +lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==>
1.1241 + (setprod f (A Un B) :: 'a ::{field})
1.1242 + = setprod f A * setprod f B / setprod f (A Int B)"
1.1243 +by (subst setprod_Un_Int [symmetric], auto)
1.1244 +
1.1245 +lemma setprod_nonneg [rule_format]:
1.1246 + "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A"
1.1247 +by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg)
1.1248 +
1.1249 +lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x)
1.1250 + --> 0 < setprod f A"
1.1251 +by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos)
1.1252 +
1.1253 +lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==>
1.1254 + (setprod f (A - {a}) :: 'a :: {field}) =
1.1255 + (if a:A then setprod f A / f a else setprod f A)"
1.1256 + by (erule finite_induct) (auto simp add: insert_Diff_if)
1.1257 +
1.1258 +lemma setprod_inversef:
1.1259 + fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero"
1.1260 + shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)"
1.1261 +by (erule finite_induct) auto
1.1262 +
1.1263 +lemma setprod_dividef:
1.1264 + fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero"
1.1265 + shows "finite A
1.1266 + ==> setprod (%x. f x / g x) A = setprod f A / setprod g A"
1.1267 +apply (subgoal_tac
1.1268 + "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A")
1.1269 +apply (erule ssubst)
1.1270 +apply (subst divide_inverse)
1.1271 +apply (subst setprod_timesf)
1.1272 +apply (subst setprod_inversef, assumption+, rule refl)
1.1273 +apply (rule setprod_cong, rule refl)
1.1274 +apply (subst divide_inverse, auto)
1.1275 +done
1.1276 +
1.1277 +lemma setprod_dvd_setprod [rule_format]:
1.1278 + "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A"
1.1279 + apply (cases "finite A")
1.1280 + apply (induct set: finite)
1.1281 + apply (auto simp add: dvd_def)
1.1282 + apply (rule_tac x = "k * ka" in exI)
1.1283 + apply (simp add: algebra_simps)
1.1284 +done
1.1285 +
1.1286 +lemma setprod_dvd_setprod_subset:
1.1287 + "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B"
1.1288 + apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)")
1.1289 + apply (unfold dvd_def, blast)
1.1290 + apply (subst setprod_Un_disjoint [symmetric])
1.1291 + apply (auto elim: finite_subset intro: setprod_cong)
1.1292 +done
1.1293 +
1.1294 +lemma setprod_dvd_setprod_subset2:
1.1295 + "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow>
1.1296 + setprod f A dvd setprod g B"
1.1297 + apply (rule dvd_trans)
1.1298 + apply (rule setprod_dvd_setprod, erule (1) bspec)
1.1299 + apply (erule (1) setprod_dvd_setprod_subset)
1.1300 +done
1.1301 +
1.1302 +lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow>
1.1303 + (f i ::'a::comm_semiring_1) dvd setprod f A"
1.1304 +by (induct set: finite) (auto intro: dvd_mult)
1.1305 +
1.1306 +lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow>
1.1307 + (d::'a::comm_semiring_1) dvd (SUM x : A. f x)"
1.1308 + apply (cases "finite A")
1.1309 + apply (induct set: finite)
1.1310 + apply auto
1.1311 +done
1.1312 +
1.1313 +lemma setprod_mono:
1.1314 + fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom"
1.1315 + assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i"
1.1316 + shows "setprod f A \<le> setprod g A"
1.1317 +proof (cases "finite A")
1.1318 + case True
1.1319 + hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A]
1.1320 + proof (induct A rule: finite_subset_induct)
1.1321 + case (insert a F)
1.1322 + thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)"
1.1323 + unfolding setprod_insert[OF insert(1,3)]
1.1324 + using assms[rule_format,OF insert(2)] insert
1.1325 + by (auto intro: mult_mono mult_nonneg_nonneg)
1.1326 + qed auto
1.1327 + thus ?thesis by simp
1.1328 +qed auto
1.1329 +
1.1330 +lemma abs_setprod:
1.1331 + fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
1.1332 + shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A"
1.1333 +proof (cases "finite A")
1.1334 + case True thus ?thesis
1.1335 + by induct (auto simp add: field_simps abs_mult)
1.1336 +qed auto
1.1337 +
1.1338 +lemma setprod_constant: "finite A ==> (\<Prod>x\<in> A. (y::'a::{comm_monoid_mult})) = y^(card A)"
1.1339 +apply (erule finite_induct)
1.1340 +apply auto
1.1341 +done
1.1342 +
1.1343 +lemma setprod_gen_delta:
1.1344 + assumes fS: "finite S"
1.1345 + 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)"
1.1346 +proof-
1.1347 + let ?f = "(\<lambda>k. if k=a then b k else c)"
1.1348 + {assume a: "a \<notin> S"
1.1349 + hence "\<forall> k\<in> S. ?f k = c" by simp
1.1350 + hence ?thesis using a setprod_constant[OF fS, of c] by simp }
1.1351 + moreover
1.1352 + {assume a: "a \<in> S"
1.1353 + let ?A = "S - {a}"
1.1354 + let ?B = "{a}"
1.1355 + have eq: "S = ?A \<union> ?B" using a by blast
1.1356 + have dj: "?A \<inter> ?B = {}" by simp
1.1357 + from fS have fAB: "finite ?A" "finite ?B" by auto
1.1358 + have fA0:"setprod ?f ?A = setprod (\<lambda>i. c) ?A"
1.1359 + apply (rule setprod_cong) by auto
1.1360 + have cA: "card ?A = card S - 1" using fS a by auto
1.1361 + have fA1: "setprod ?f ?A = c ^ card ?A" unfolding fA0 apply (rule setprod_constant) using fS by auto
1.1362 + have "setprod ?f ?A * setprod ?f ?B = setprod ?f S"
1.1363 + using setprod_Un_disjoint[OF fAB dj, of ?f, unfolded eq[symmetric]]
1.1364 + by simp
1.1365 + then have ?thesis using a cA
1.1366 + by (simp add: fA1 field_simps cong add: setprod_cong cong del: if_weak_cong)}
1.1367 + ultimately show ?thesis by blast
1.1368 +qed
1.1369 +
1.1370 +lemma setprod_eq_1_iff [simp]:
1.1371 + "finite F ==> setprod f F = 1 \<longleftrightarrow> (ALL a:F. f a = (1::nat))"
1.1372 + by (induct set: finite) auto
1.1373 +
1.1374 +lemma setprod_pos_nat:
1.1375 + "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0"
1.1376 +using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
1.1377 +
1.1378 +lemma setprod_pos_nat_iff[simp]:
1.1379 + "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))"
1.1380 +using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric])
1.1381 +
1.1382 +end