| author | blanchet | 
| Wed, 12 Feb 2014 17:35:59 +0100 | |
| changeset 55443 | 3def821deb70 | 
| parent 55096 | 916b2ac758f4 | 
| child 56166 | 9a241bc276cd | 
| permissions | -rw-r--r-- | 
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1 | (* Title: HOL/Groups_Big.thy | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 2 | Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 3 | with contributions by Jeremy Avigad | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 4 | *) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 5 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 6 | header {* Big sum and product over finite (non-empty) sets *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 7 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 8 | theory Groups_Big | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 9 | imports Finite_Set | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 10 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 11 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 12 | subsection {* Generic monoid operation over a set *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 13 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 14 | no_notation times (infixl "*" 70) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 15 | no_notation Groups.one ("1")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 16 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 17 | locale comm_monoid_set = comm_monoid | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 18 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 19 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 20 | interpretation comp_fun_commute f | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 21 | by default (simp add: fun_eq_iff left_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 22 | |
| 54745 | 23 | interpretation comp?: comp_fun_commute "f \<circ> g" | 
| 24 | by (fact comp_comp_fun_commute) | |
| 54744 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 25 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 26 | definition F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 27 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 28 | eq_fold: "F g A = Finite_Set.fold (f \<circ> g) 1 A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 29 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 30 | lemma infinite [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 31 | "\<not> finite A \<Longrightarrow> F g A = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 32 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 33 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 34 | lemma empty [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 35 |   "F g {} = 1"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 36 | by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 37 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 38 | lemma insert [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 39 | assumes "finite A" and "x \<notin> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 40 | shows "F g (insert x A) = g x * F g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 41 | using assms by (simp add: eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 42 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 43 | lemma remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 44 | assumes "finite A" and "x \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 45 |   shows "F g A = g x * F g (A - {x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 46 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 47 | from `x \<in> A` obtain B where A: "A = insert x B" and "x \<notin> B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 48 | by (auto dest: mk_disjoint_insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 49 | moreover from `finite A` A have "finite B" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 50 | ultimately show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 51 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 52 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 53 | lemma insert_remove: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 54 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 55 |   shows "F g (insert x A) = g x * F g (A - {x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 56 | using assms by (cases "x \<in> A") (simp_all add: remove insert_absorb) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 57 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 58 | lemma neutral: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 59 | assumes "\<forall>x\<in>A. g x = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 60 | shows "F g A = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 61 | using assms by (induct A rule: infinite_finite_induct) simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 62 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 63 | lemma neutral_const [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 64 | "F (\<lambda>_. 1) A = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 65 | by (simp add: neutral) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 66 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 67 | lemma union_inter: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 68 | assumes "finite A" and "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 69 | shows "F g (A \<union> B) * F g (A \<inter> B) = F g A * F g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 70 |   -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 71 | using assms proof (induct A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 72 | case empty then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 73 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 74 | case (insert x A) then show ?case | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 75 | by (auto simp add: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 76 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 77 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 78 | corollary union_inter_neutral: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 79 | assumes "finite A" and "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 80 | and I0: "\<forall>x \<in> A \<inter> B. g x = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 81 | shows "F g (A \<union> B) = F g A * F g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 82 | using assms by (simp add: union_inter [symmetric] neutral) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 83 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 84 | corollary union_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 85 | assumes "finite A" and "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 86 |   assumes "A \<inter> B = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 87 | shows "F g (A \<union> B) = F g A * F g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 88 | using assms by (simp add: union_inter_neutral) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 89 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 90 | lemma subset_diff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 91 | assumes "B \<subseteq> A" and "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 92 | shows "F g A = F g (A - B) * F g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 93 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 94 | from assms have "finite (A - B)" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 95 | moreover from assms have "finite B" by (rule finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 96 |   moreover from assms have "(A - B) \<inter> B = {}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 97 | ultimately have "F g (A - B \<union> B) = F g (A - B) * F g B" by (rule union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 98 | moreover from assms have "A \<union> B = A" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 99 | ultimately show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 100 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 101 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 102 | lemma reindex: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 103 | assumes "inj_on h A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 104 | shows "F g (h ` A) = F (g \<circ> h) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 105 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 106 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 107 | with assms show ?thesis by (simp add: eq_fold fold_image comp_assoc) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 108 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 109 | case False with assms have "\<not> finite (h ` A)" by (blast dest: finite_imageD) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 110 | with False show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 111 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 112 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 113 | lemma cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 114 | assumes "A = B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 115 | assumes g_h: "\<And>x. x \<in> B \<Longrightarrow> g x = h x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 116 | shows "F g A = F h B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 117 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 118 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 119 | then have "\<And>C. C \<subseteq> A \<longrightarrow> (\<forall>x\<in>C. g x = h x) \<longrightarrow> F g C = F h C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 120 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 121 | case empty then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 122 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 123 | case (insert x F) then show ?case apply - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 124 | apply (simp add: subset_insert_iff, clarify) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 125 | apply (subgoal_tac "finite C") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 126 | prefer 2 apply (blast dest: finite_subset [rotated]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 127 |     apply (subgoal_tac "C = insert x (C - {x})")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 128 | prefer 2 apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 129 | apply (erule ssubst) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 130 | apply (simp add: Ball_def del: insert_Diff_single) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 131 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 132 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 133 | with `A = B` g_h show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 134 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 135 | case False | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 136 | with `A = B` show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 137 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 138 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 139 | lemma strong_cong [cong]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 140 | assumes "A = B" "\<And>x. x \<in> B =simp=> g x = h x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 141 | shows "F (\<lambda>x. g x) A = F (\<lambda>x. h x) B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 142 | by (rule cong) (insert assms, simp_all add: simp_implies_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 143 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 144 | lemma UNION_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 145 | assumes "finite I" and "\<forall>i\<in>I. finite (A i)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 146 |   and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 147 | shows "F g (UNION I A) = F (\<lambda>x. F g (A x)) I" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 148 | apply (insert assms) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 149 | apply (induct rule: finite_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 150 | apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 151 | apply atomize | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 152 | apply (subgoal_tac "\<forall>i\<in>Fa. x \<noteq> i") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 153 | prefer 2 apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 154 | apply (subgoal_tac "A x Int UNION Fa A = {}")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 155 | prefer 2 apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 156 | apply (simp add: union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 157 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 158 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 159 | lemma Union_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 160 |   assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 161 | shows "F g (Union C) = F (F g) C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 162 | proof cases | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 163 | assume "finite C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 164 | from UNION_disjoint [OF this assms] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 165 | show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 166 | by (simp add: SUP_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 167 | qed (auto dest: finite_UnionD intro: infinite) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 168 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 169 | lemma distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 170 | "F (\<lambda>x. g x * h x) A = F g A * F h A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 171 | using assms by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 172 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 173 | lemma Sigma: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 174 | "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)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 175 | apply (subst Sigma_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 176 | apply (subst UNION_disjoint, assumption, simp) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 177 | apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 178 | apply (rule cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 179 | apply rule | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 180 | apply (simp add: fun_eq_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 181 | apply (subst UNION_disjoint, simp, simp) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 182 | apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 183 | apply (simp add: comp_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 184 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 185 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 186 | lemma related: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 187 | assumes Re: "R 1 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 188 | and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 * y1) (x2 * y2)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 189 | and fS: "finite S" and Rfg: "\<forall>x\<in>S. R (h x) (g x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 190 | shows "R (F h S) (F g S)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 191 | using fS by (rule finite_subset_induct) (insert assms, auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 192 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 193 | lemma eq_general: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 194 | assumes h: "\<forall>y\<in>S'. \<exists>!x. x \<in> S \<and> h x = y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 195 | and f12: "\<forall>x\<in>S. h x \<in> S' \<and> f2 (h x) = f1 x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 196 | shows "F f1 S = F f2 S'" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 197 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 198 | from h f12 have hS: "h ` S = S'" by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 199 |   {fix x y assume H: "x \<in> S" "y \<in> S" "h x = h y"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 200 | from f12 h H have "x = y" by auto } | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 201 | hence hinj: "inj_on h S" unfolding inj_on_def Ex1_def by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 202 | from f12 have th: "\<And>x. x \<in> S \<Longrightarrow> (f2 \<circ> h) x = f1 x" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 203 | from hS have "F f2 S' = F f2 (h ` S)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 204 | also have "\<dots> = F (f2 o h) S" using reindex [OF hinj, of f2] . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 205 | also have "\<dots> = F f1 S " using th cong [of _ _ "f2 o h" f1] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 206 | by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 207 | finally show ?thesis .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 208 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 209 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 210 | lemma eq_general_reverses: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 211 | assumes kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 212 | and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = j x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 213 | shows "F j S = F g T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 214 | (* metis solves it, but not yet available here *) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 215 | apply (rule eq_general [of T S h g j]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 216 | apply (rule ballI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 217 | apply (frule kh) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 218 | apply (rule ex1I[]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 219 | apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 220 | apply clarsimp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 221 | apply (drule hk) apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 222 | apply (rule sym) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 223 | apply (erule conjunct1[OF conjunct2[OF hk]]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 224 | apply (rule ballI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 225 | apply (drule hk) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 226 | apply blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 227 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 228 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 229 | lemma mono_neutral_cong_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 230 | assumes "finite T" and "S \<subseteq> T" and "\<forall>i \<in> T - S. h i = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 231 | and "\<And>x. x \<in> S \<Longrightarrow> g x = h x" shows "F g S = F h T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 232 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 233 | have eq: "T = S \<union> (T - S)" using `S \<subseteq> T` by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 234 |   have d: "S \<inter> (T - S) = {}" using `S \<subseteq> T` by blast
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 235 | from `finite T` `S \<subseteq> T` have f: "finite S" "finite (T - S)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 236 | by (auto intro: finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 237 | show ?thesis using assms(4) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 238 | by (simp add: union_disjoint [OF f d, unfolded eq [symmetric]] neutral [OF assms(3)]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 239 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 240 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 241 | lemma mono_neutral_cong_right: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 242 | "\<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> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 243 | \<Longrightarrow> F g T = F h S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 244 | by (auto intro!: mono_neutral_cong_left [symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 245 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 246 | lemma mono_neutral_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 247 | "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g S = F g T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 248 | by (blast intro: mono_neutral_cong_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 249 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 250 | lemma mono_neutral_right: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 251 | "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. g i = 1 \<rbrakk> \<Longrightarrow> F g T = F g S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 252 | by (blast intro!: mono_neutral_left [symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 253 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 254 | lemma delta: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 255 | assumes fS: "finite S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 256 | shows "F (\<lambda>k. if k = a then b k else 1) S = (if a \<in> S then b a else 1)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 257 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 258 | let ?f = "(\<lambda>k. if k=a then b k else 1)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 259 |   { assume a: "a \<notin> S"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 260 | hence "\<forall>k\<in>S. ?f k = 1" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 261 | hence ?thesis using a by simp } | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 262 | moreover | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 263 |   { assume a: "a \<in> S"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 264 |     let ?A = "S - {a}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 265 |     let ?B = "{a}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 266 | have eq: "S = ?A \<union> ?B" using a by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 267 |     have dj: "?A \<inter> ?B = {}" by simp
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 268 | from fS have fAB: "finite ?A" "finite ?B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 269 | have "F ?f S = F ?f ?A * F ?f ?B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 270 | using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 271 | by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 272 | then have ?thesis using a by simp } | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 273 | ultimately show ?thesis by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 274 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 275 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 276 | lemma delta': | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 277 | assumes fS: "finite S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 278 | shows "F (\<lambda>k. if a = k then b k else 1) S = (if a \<in> S then b a else 1)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 279 | using delta [OF fS, of a b, symmetric] by (auto intro: cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 280 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 281 | lemma If_cases: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 282 | fixes P :: "'b \<Rightarrow> bool" and g h :: "'b \<Rightarrow> 'a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 283 | assumes fA: "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 284 | shows "F (\<lambda>x. if P x then h x else g x) A = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 285 |     F h (A \<inter> {x. P x}) * F g (A \<inter> - {x. P x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 286 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 287 |   have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 288 |           "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 289 | by blast+ | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 290 | from fA | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 291 |   have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 292 | let ?g = "\<lambda>x. if P x then h x else g x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 293 | from union_disjoint [OF f a(2), of ?g] a(1) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 294 | show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 295 | by (subst (1 2) cong) simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 296 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 297 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 298 | lemma cartesian_product: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 299 | "F (\<lambda>x. F (g x) B) A = F (split g) (A <*> B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 300 | apply (rule sym) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 301 | apply (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 302 | apply (cases "finite B") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 303 | apply (simp add: Sigma) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 304 |  apply (cases "A={}", simp)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 305 | apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 306 | apply (auto intro: infinite dest: finite_cartesian_productD2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 307 | apply (cases "B = {}") apply (auto intro: infinite dest: finite_cartesian_productD1)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 308 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 309 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 310 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 311 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 312 | notation times (infixl "*" 70) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 313 | notation Groups.one ("1")
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 314 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 315 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 316 | subsection {* Generalized summation over a set *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 317 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 318 | context comm_monoid_add | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 319 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 320 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 321 | definition setsum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 322 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 323 | "setsum = comm_monoid_set.F plus 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 324 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 325 | sublocale setsum!: comm_monoid_set plus 0 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 326 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 327 | "comm_monoid_set.F plus 0 = setsum" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 328 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 329 | show "comm_monoid_set plus 0" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 330 | then interpret setsum!: comm_monoid_set plus 0 . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 331 | from setsum_def show "comm_monoid_set.F plus 0 = setsum" by rule | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 332 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 333 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 334 | abbreviation | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 335 |   Setsum ("\<Sum>_" [1000] 999) where
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 336 | "\<Sum>A \<equiv> setsum (%x. x) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 337 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 338 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 339 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 340 | text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 341 | written @{text"\<Sum>x\<in>A. e"}. *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 342 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 343 | syntax | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 344 |   "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3SUM _:_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 345 | syntax (xsymbols) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 346 |   "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 347 | syntax (HTML output) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 348 |   "_setsum" :: "pttrn => 'a set => 'b => 'b::comm_monoid_add"    ("(3\<Sum>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 349 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 350 | translations -- {* Beware of argument permutation! *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 351 | "SUM i:A. b" == "CONST setsum (%i. b) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 352 | "\<Sum>i\<in>A. b" == "CONST setsum (%i. b) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 353 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 354 | text{* Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 355 |  @{text"\<Sum>x|P. e"}. *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 356 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 357 | syntax | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 358 |   "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3SUM _ |/ _./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 359 | syntax (xsymbols) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 360 |   "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 361 | syntax (HTML output) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 362 |   "_qsetsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Sum>_ | (_)./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 363 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 364 | translations | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 365 |   "SUM x|P. t" => "CONST setsum (%x. t) {x. P}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 366 |   "\<Sum>x|P. t" => "CONST setsum (%x. t) {x. P}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 367 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 368 | print_translation {*
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 369 | let | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 370 |   fun setsum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 371 | if x <> y then raise Match | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 372 | else | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 373 | let | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 374 | val x' = Syntax_Trans.mark_bound_body (x, Tx); | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 375 | val t' = subst_bound (x', t); | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 376 | val P' = subst_bound (x', P); | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 377 | in | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 378 |             Syntax.const @{syntax_const "_qsetsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 379 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 380 | | setsum_tr' _ = raise Match; | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 381 | in [(@{const_syntax setsum}, K setsum_tr')] end
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 382 | *} | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 383 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 384 | text {* TODO These are candidates for generalization *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 385 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 386 | context comm_monoid_add | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 387 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 388 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 389 | lemma setsum_reindex_id: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 390 | "inj_on f B ==> setsum f B = setsum id (f ` B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 391 | by (simp add: setsum.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 392 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 393 | lemma setsum_reindex_nonzero: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 394 | assumes fS: "finite S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 395 | 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" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 396 | shows "setsum h (f ` S) = setsum (h \<circ> f) S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 397 | using nz proof (induct rule: finite_induct [OF fS]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 398 | case 1 thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 399 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 400 | case (2 x F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 401 |   { assume fxF: "f x \<in> f ` F" hence "\<exists>y \<in> F . f y = f x" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 402 | then obtain y where y: "y \<in> F" "f x = f y" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 403 | from "2.hyps" y have xy: "x \<noteq> y" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 404 | from "2.prems" [of x y] "2.hyps" xy y have h0: "h (f x) = 0" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 405 | have "setsum h (f ` insert x F) = setsum h (f ` F)" using fxF by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 406 | also have "\<dots> = setsum (h o f) (insert x F)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 407 | unfolding setsum.insert[OF `finite F` `x\<notin>F`] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 408 | using h0 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 409 | apply (simp cong del: setsum.strong_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 410 | apply (rule "2.hyps"(3)) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 411 | apply (rule_tac y="y" in "2.prems") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 412 | apply simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 413 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 414 | finally have ?case . } | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 415 | moreover | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 416 |   { assume fxF: "f x \<notin> f ` F"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 417 | have "setsum h (f ` insert x F) = h (f x) + setsum h (f ` F)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 418 | using fxF "2.hyps" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 419 | also have "\<dots> = setsum (h o f) (insert x F)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 420 | unfolding setsum.insert[OF `finite F` `x\<notin>F`] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 421 | apply (simp cong del: setsum.strong_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 422 | apply (rule cong [OF refl [of "op + (h (f x))"]]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 423 | apply (rule "2.hyps"(3)) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 424 | apply (rule_tac y="y" in "2.prems") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 425 | apply simp_all | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 426 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 427 | finally have ?case . } | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 428 | ultimately show ?case by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 429 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 430 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 431 | lemma setsum_cong2: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 432 | "(\<And>x. x \<in> A \<Longrightarrow> f x = g x) \<Longrightarrow> setsum f A = setsum g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 433 | by (auto intro: setsum.cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 434 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 435 | lemma setsum_reindex_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 436 | "[|inj_on f A; B = f ` A; !!a. a:A \<Longrightarrow> g a = h (f a)|] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 437 | ==> setsum h B = setsum g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 438 | by (simp add: setsum.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 439 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 440 | lemma setsum_restrict_set: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 441 | assumes fA: "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 442 | shows "setsum f (A \<inter> B) = setsum (\<lambda>x. if x \<in> B then f x else 0) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 443 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 444 | from fA have fab: "finite (A \<inter> B)" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 445 | have aba: "A \<inter> B \<subseteq> A" by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 446 | let ?g = "\<lambda>x. if x \<in> A\<inter>B then f x else 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 447 | from setsum.mono_neutral_left [OF fA aba, of ?g] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 448 | show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 449 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 450 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 451 | lemma setsum_Union_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 452 |   assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 453 | shows "setsum f (Union C) = setsum (setsum f) C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 454 | using assms by (fact setsum.Union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 455 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 456 | lemma setsum_cartesian_product: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 457 | "(\<Sum>x\<in>A. (\<Sum>y\<in>B. f x y)) = (\<Sum>(x,y) \<in> A <*> B. f x y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 458 | by (fact setsum.cartesian_product) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 459 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 460 | lemma setsum_UNION_zero: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 461 | assumes fS: "finite S" and fSS: "\<forall>T \<in> S. finite T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 462 | 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" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 463 | shows "setsum f (\<Union>S) = setsum (\<lambda>T. setsum f T) S" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 464 | using fSS f0 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 465 | proof(induct rule: finite_induct[OF fS]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 466 | case 1 thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 467 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 468 | case (2 T F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 469 | then have fTF: "finite T" "\<forall>T\<in>F. finite T" "finite F" and TF: "T \<notin> F" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 470 | and H: "setsum f (\<Union> F) = setsum (setsum f) F" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 471 | from fTF have fUF: "finite (\<Union>F)" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 472 | from "2.prems" TF fTF | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 473 | show ?case | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 474 | by (auto simp add: H [symmetric] intro: setsum.union_inter_neutral [OF fTF(1) fUF, of f]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 475 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 476 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 477 | text {* Commuting outer and inner summation *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 478 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 479 | lemma setsum_commute: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 480 | "(\<Sum>i\<in>A. \<Sum>j\<in>B. f i j) = (\<Sum>j\<in>B. \<Sum>i\<in>A. f i j)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 481 | proof (simp add: setsum_cartesian_product) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 482 | have "(\<Sum>(x,y) \<in> A <*> B. f x y) = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 483 | (\<Sum>(y,x) \<in> (%(i, j). (j, i)) ` (A \<times> B). f x y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 484 | (is "?s = _") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 485 | apply (simp add: setsum.reindex [where h = "%(i, j). (j, i)"] swap_inj_on) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 486 | apply (simp add: split_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 487 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 488 | also have "... = (\<Sum>(y,x)\<in>B \<times> A. f x y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 489 | (is "_ = ?t") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 490 | apply (simp add: swap_product) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 491 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 492 | finally show "?s = ?t" . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 493 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 494 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 495 | lemma setsum_Plus: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 496 | fixes A :: "'a set" and B :: "'b set" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 497 | assumes fin: "finite A" "finite B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 498 | shows "setsum f (A <+> B) = setsum (f \<circ> Inl) A + setsum (f \<circ> Inr) B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 499 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 500 | have "A <+> B = Inl ` A \<union> Inr ` B" by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 501 |   moreover from fin have "finite (Inl ` A :: ('a + 'b) set)" "finite (Inr ` B :: ('a + 'b) set)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 502 | by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 503 |   moreover have "Inl ` A \<inter> Inr ` B = ({} :: ('a + 'b) set)" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 504 | moreover have "inj_on (Inl :: 'a \<Rightarrow> 'a + 'b) A" "inj_on (Inr :: 'b \<Rightarrow> 'a + 'b) B" by(auto intro: inj_onI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 505 | ultimately show ?thesis using fin by(simp add: setsum.union_disjoint setsum.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 506 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 507 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 508 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 509 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 510 | text {* TODO These are legacy *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 511 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 512 | lemma setsum_empty: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 513 |   "setsum f {} = 0"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 514 | by (fact setsum.empty) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 515 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 516 | lemma setsum_insert: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 517 | "finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 518 | by (fact setsum.insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 519 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 520 | lemma setsum_infinite: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 521 | "~ finite A ==> setsum f A = 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 522 | by (fact setsum.infinite) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 523 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 524 | lemma setsum_reindex: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 525 | "inj_on f B \<Longrightarrow> setsum h (f ` B) = setsum (h \<circ> f) B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 526 | by (fact setsum.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 527 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 528 | lemma setsum_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 529 | "A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 530 | by (fact setsum.cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 531 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 532 | lemma strong_setsum_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 533 | "A = B ==> (!!x. x:B =simp=> f x = g x) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 534 | ==> setsum (%x. f x) A = setsum (%x. g x) B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 535 | by (fact setsum.strong_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 536 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 537 | lemmas setsum_0 = setsum.neutral_const | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 538 | lemmas setsum_0' = setsum.neutral | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 539 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 540 | lemma setsum_Un_Int: "finite A ==> finite B ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 541 | setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 542 |   -- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 543 | by (fact setsum.union_inter) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 544 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 545 | lemma setsum_Un_disjoint: "finite A ==> finite B | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 546 |   ==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 547 | by (fact setsum.union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 548 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 549 | lemma setsum_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 550 | setsum f A = setsum f (A - B) + setsum f B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 551 | by (fact setsum.subset_diff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 552 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 553 | lemma setsum_mono_zero_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 554 | "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 0 \<rbrakk> \<Longrightarrow> setsum f S = setsum f T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 555 | by (fact setsum.mono_neutral_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 556 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 557 | lemmas setsum_mono_zero_right = setsum.mono_neutral_right | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 558 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 559 | lemma setsum_mono_zero_cong_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 560 | "\<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> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 561 | \<Longrightarrow> setsum f S = setsum g T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 562 | by (fact setsum.mono_neutral_cong_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 563 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 564 | lemmas setsum_mono_zero_cong_right = setsum.mono_neutral_cong_right | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 565 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 566 | lemma setsum_delta: "finite S \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 567 | setsum (\<lambda>k. if k=a then b k else 0) S = (if a \<in> S then b a else 0)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 568 | by (fact setsum.delta) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 569 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 570 | lemma setsum_delta': "finite S \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 571 | setsum (\<lambda>k. if a = k then b k else 0) S = (if a\<in> S then b a else 0)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 572 | by (fact setsum.delta') | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 573 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 574 | lemma setsum_cases: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 575 | assumes "finite A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 576 | shows "setsum (\<lambda>x. if P x then f x else g x) A = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 577 |          setsum f (A \<inter> {x. P x}) + setsum g (A \<inter> - {x. P x})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 578 | using assms by (fact setsum.If_cases) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 579 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 580 | (*But we can't get rid of finite I. If infinite, although the rhs is 0, | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 581 | the lhs need not be, since UNION I A could still be finite.*) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 582 | lemma setsum_UN_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 583 | assumes "finite I" and "ALL i:I. finite (A i)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 584 |     and "ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 585 | shows "setsum f (UNION I A) = (\<Sum>i\<in>I. setsum f (A i))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 586 | using assms by (fact setsum.UNION_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 587 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 588 | (*But we can't get rid of finite A. If infinite, although the lhs is 0, | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 589 | the rhs need not be, since SIGMA A B could still be finite.*) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 590 | lemma setsum_Sigma: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 591 | assumes "finite A" and "ALL x:A. finite (B x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 592 | 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)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 593 | using assms by (fact setsum.Sigma) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 594 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 595 | lemma setsum_addf: "setsum (%x. f x + g x) A = (setsum f A + setsum g A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 596 | by (fact setsum.distrib) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 597 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 598 | lemma setsum_Un_zero: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 599 | "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 0 \<rbrakk> \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 600 | setsum f (S \<union> T) = setsum f S + setsum f T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 601 | by (fact setsum.union_inter_neutral) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 602 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 603 | lemma setsum_eq_general_reverses: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 604 | assumes fS: "finite S" and fT: "finite T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 605 | and kh: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S \<and> h (k y) = y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 606 | and hk: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> T \<and> k (h x) = x \<and> g (h x) = f x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 607 | shows "setsum f S = setsum g T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 608 | using kh hk by (fact setsum.eq_general_reverses) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 609 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 610 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 611 | subsubsection {* Properties in more restricted classes of structures *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 612 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 613 | lemma setsum_Un: "finite A ==> finite B ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 614 | (setsum f (A Un B) :: 'a :: ab_group_add) = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 615 | setsum f A + setsum f B - setsum f (A Int B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 616 | by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 617 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 618 | lemma setsum_Un2: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 619 | assumes "finite (A \<union> B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 620 | shows "setsum f (A \<union> B) = setsum f (A - B) + setsum f (B - A) + setsum f (A \<inter> B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 621 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 622 | have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 623 | by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 624 | with assms show ?thesis by simp (subst setsum_Un_disjoint, auto)+ | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 625 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 626 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 627 | lemma setsum_diff1: "finite A \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 628 |   (setsum f (A - {a}) :: ('a::ab_group_add)) =
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 629 | (if a:A then setsum f A - f a else setsum f A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 630 | by (erule finite_induct) (auto simp add: insert_Diff_if) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 631 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 632 | lemma setsum_diff: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 633 | assumes le: "finite A" "B \<subseteq> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 634 |   shows "setsum f (A - B) = setsum f A - ((setsum f B)::('a::ab_group_add))"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 635 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 636 | from le have finiteB: "finite B" using finite_subset by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 637 | show ?thesis using finiteB le | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 638 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 639 | case empty | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 640 | thus ?case by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 641 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 642 | case (insert x F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 643 | thus ?case using le finiteB | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 644 | by (simp add: Diff_insert[where a=x and B=F] setsum_diff1 insert_absorb) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 645 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 646 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 647 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 648 | lemma setsum_mono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 649 |   assumes le: "\<And>i. i\<in>K \<Longrightarrow> f (i::'a) \<le> ((g i)::('b::{comm_monoid_add, ordered_ab_semigroup_add}))"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 650 | shows "(\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 651 | proof (cases "finite K") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 652 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 653 | thus ?thesis using le | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 654 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 655 | case empty | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 656 | thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 657 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 658 | case insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 659 | thus ?case using add_mono by fastforce | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 660 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 661 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 662 | case False then show ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 663 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 664 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 665 | lemma setsum_strict_mono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 666 |   fixes f :: "'a \<Rightarrow> 'b::{ordered_cancel_ab_semigroup_add,comm_monoid_add}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 667 |   assumes "finite A"  "A \<noteq> {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 668 | and "!!x. x:A \<Longrightarrow> f x < g x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 669 | shows "setsum f A < setsum g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 670 | using assms | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 671 | proof (induct rule: finite_ne_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 672 | case singleton thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 673 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 674 | case insert thus ?case by (auto simp: add_strict_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 675 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 676 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 677 | lemma setsum_strict_mono_ex1: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 678 | fixes f :: "'a \<Rightarrow> 'b::{comm_monoid_add, ordered_cancel_ab_semigroup_add}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 679 | assumes "finite A" and "ALL x:A. f x \<le> g x" and "EX a:A. f a < g a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 680 | shows "setsum f A < setsum g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 681 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 682 | from assms(3) obtain a where a: "a:A" "f a < g a" by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 683 |   have "setsum f A = setsum f ((A-{a}) \<union> {a})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 684 | by(simp add:insert_absorb[OF `a:A`]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 685 |   also have "\<dots> = setsum f (A-{a}) + setsum f {a}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 686 | using `finite A` by(subst setsum_Un_disjoint) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 687 |   also have "setsum f (A-{a}) \<le> setsum g (A-{a})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 688 | by(rule setsum_mono)(simp add: assms(2)) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 689 |   also have "setsum f {a} < setsum g {a}" using a by simp
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 690 |   also have "setsum g (A - {a}) + setsum g {a} = setsum g((A-{a}) \<union> {a})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 691 | using `finite A` by(subst setsum_Un_disjoint[symmetric]) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 692 | also have "\<dots> = setsum g A" by(simp add:insert_absorb[OF `a:A`]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 693 | finally show ?thesis by (auto simp add: add_right_mono add_strict_left_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 694 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 695 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 696 | lemma setsum_negf: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 697 | "setsum (%x. - (f x)::'a::ab_group_add) A = - setsum f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 698 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 699 | case True thus ?thesis by (induct set: finite) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 700 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 701 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 702 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 703 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 704 | lemma setsum_subtractf: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 705 | "setsum (%x. ((f x)::'a::ab_group_add) - g x) A = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 706 | setsum f A - setsum g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 707 | using setsum_addf [of f "- g" A] by (simp add: setsum_negf) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 708 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 709 | lemma setsum_nonneg: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 710 |   assumes nn: "\<forall>x\<in>A. (0::'a::{ordered_ab_semigroup_add,comm_monoid_add}) \<le> f x"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 711 | shows "0 \<le> setsum f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 712 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 713 | case True thus ?thesis using nn | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 714 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 715 | case empty then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 716 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 717 | case (insert x F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 718 | then have "0 + 0 \<le> f x + setsum f F" by (blast intro: add_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 719 | with insert show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 720 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 721 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 722 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 723 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 724 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 725 | lemma setsum_nonpos: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 726 |   assumes np: "\<forall>x\<in>A. f x \<le> (0::'a::{ordered_ab_semigroup_add,comm_monoid_add})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 727 | shows "setsum f A \<le> 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 728 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 729 | case True thus ?thesis using np | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 730 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 731 | case empty then show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 732 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 733 | case (insert x F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 734 | then have "f x + setsum f F \<le> 0 + 0" by (blast intro: add_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 735 | with insert show ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 736 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 737 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 738 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 739 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 740 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 741 | lemma setsum_nonneg_leq_bound: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 742 |   fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 743 | assumes "finite s" "\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0" "(\<Sum>i \<in> s. f i) = B" "i \<in> s" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 744 | shows "f i \<le> B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 745 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 746 |   have "0 \<le> (\<Sum> i \<in> s - {i}. f i)" and "0 \<le> f i"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 747 | using assms by (auto intro!: setsum_nonneg) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 748 | moreover | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 749 |   have "(\<Sum> i \<in> s - {i}. f i) + f i = B"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 750 | using assms by (simp add: setsum_diff1) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 751 | ultimately show ?thesis by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 752 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 753 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 754 | lemma setsum_nonneg_0: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 755 |   fixes f :: "'a \<Rightarrow> 'b::{ordered_ab_group_add}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 756 | assumes "finite s" and pos: "\<And> i. i \<in> s \<Longrightarrow> f i \<ge> 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 757 | and "(\<Sum> i \<in> s. f i) = 0" and i: "i \<in> s" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 758 | shows "f i = 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 759 | using setsum_nonneg_leq_bound[OF assms] pos[OF i] by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 760 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 761 | lemma setsum_mono2: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 762 | fixes f :: "'a \<Rightarrow> 'b :: ordered_comm_monoid_add" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 763 | assumes fin: "finite B" and sub: "A \<subseteq> B" and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 764 | shows "setsum f A \<le> setsum f B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 765 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 766 | have "setsum f A \<le> setsum f A + setsum f (B-A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 767 | by(simp add: add_increasing2[OF setsum_nonneg] nn Ball_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 768 | also have "\<dots> = setsum f (A \<union> (B-A))" using fin finite_subset[OF sub fin] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 769 | by (simp add:setsum_Un_disjoint del:Un_Diff_cancel) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 770 | also have "A \<union> (B-A) = B" using sub by blast | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 771 | finally show ?thesis . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 772 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 773 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 774 | lemma setsum_mono3: "finite B ==> A <= B ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 775 | ALL x: B - A. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 776 |       0 <= ((f x)::'a::{comm_monoid_add,ordered_ab_semigroup_add}) ==>
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 777 | setsum f A <= setsum f B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 778 | apply (subgoal_tac "setsum f B = setsum f A + setsum f (B - A)") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 779 | apply (erule ssubst) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 780 | apply (subgoal_tac "setsum f A + 0 <= setsum f A + setsum f (B - A)") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 781 | apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 782 | apply (rule add_left_mono) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 783 | apply (erule setsum_nonneg) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 784 | apply (subst setsum_Un_disjoint [THEN sym]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 785 | apply (erule finite_subset, assumption) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 786 | apply (rule finite_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 787 | prefer 2 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 788 | apply assumption | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 789 | apply (auto simp add: sup_absorb2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 790 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 791 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 792 | lemma setsum_right_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 793 |   fixes f :: "'a => ('b::semiring_0)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 794 | shows "r * setsum f A = setsum (%n. r * f n) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 795 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 796 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 797 | thus ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 798 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 799 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 800 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 801 | case (insert x A) thus ?case by (simp add: distrib_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 802 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 803 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 804 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 805 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 806 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 807 | lemma setsum_left_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 808 | "setsum f A * (r::'a::semiring_0) = (\<Sum>n\<in>A. f n * r)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 809 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 810 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 811 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 812 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 813 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 814 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 815 | case (insert x A) thus ?case by (simp add: distrib_right) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 816 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 817 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 818 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 819 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 820 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 821 | lemma setsum_divide_distrib: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 822 | "setsum f A / (r::'a::field) = (\<Sum>n\<in>A. f n / r)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 823 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 824 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 825 | then show ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 826 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 827 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 828 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 829 | case (insert x A) thus ?case by (simp add: add_divide_distrib) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 830 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 831 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 832 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 833 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 834 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 835 | lemma setsum_abs[iff]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 836 |   fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 837 | shows "abs (setsum f A) \<le> setsum (%i. abs(f i)) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 838 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 839 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 840 | thus ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 841 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 842 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 843 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 844 | case (insert x A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 845 | thus ?case by (auto intro: abs_triangle_ineq order_trans) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 846 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 847 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 848 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 849 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 850 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 851 | lemma setsum_abs_ge_zero[iff]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 852 |   fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 853 | shows "0 \<le> setsum (%i. abs(f i)) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 854 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 855 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 856 | thus ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 857 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 858 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 859 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 860 | case (insert x A) thus ?case by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 861 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 862 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 863 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 864 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 865 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 866 | lemma abs_setsum_abs[simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 867 |   fixes f :: "'a => ('b::ordered_ab_group_add_abs)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 868 | shows "abs (\<Sum>a\<in>A. abs(f a)) = (\<Sum>a\<in>A. abs(f a))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 869 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 870 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 871 | thus ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 872 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 873 | case empty thus ?case by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 874 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 875 | case (insert a A) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 876 | 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 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 877 | also have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" using insert by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 878 | also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 879 | by (simp del: abs_of_nonneg) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 880 | also have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" using insert by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 881 | finally show ?case . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 882 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 883 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 884 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 885 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 886 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 887 | lemma setsum_diff1'[rule_format]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 888 |   "finite A \<Longrightarrow> a \<in> A \<longrightarrow> (\<Sum> x \<in> A. f x) = f a + (\<Sum> x \<in> (A - {a}). f x)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 889 | 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))"])
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 890 | apply (auto simp add: insert_Diff_if add_ac) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 891 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 892 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 893 | lemma setsum_diff1_ring: assumes "finite A" "a \<in> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 894 |   shows "setsum f (A - {a}) = setsum f A - (f a::'a::ring)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 895 | unfolding setsum_diff1'[OF assms] by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 896 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 897 | lemma setsum_product: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 898 |   fixes f :: "'a => ('b::semiring_0)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 899 | shows "setsum f A * setsum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 900 | by (simp add: setsum_right_distrib setsum_left_distrib) (rule setsum_commute) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 901 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 902 | lemma setsum_mult_setsum_if_inj: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 903 | fixes f :: "'a => ('b::semiring_0)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 904 | shows "inj_on (%(a,b). f a * g b) (A \<times> B) ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 905 |   setsum f A * setsum g B = setsum id {f a * g b|a b. a:A & b:B}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 906 | by(auto simp: setsum_product setsum_cartesian_product | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 907 | intro!: setsum_reindex_cong[symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 908 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 909 | lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 910 | apply (case_tac "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 911 | prefer 2 apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 912 | apply (erule rev_mp) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 913 | apply (erule finite_induct, auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 914 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 915 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 916 | lemma setsum_eq_0_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 917 | "finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 918 | by (induct set: finite) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 919 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 920 | lemma setsum_eq_Suc0_iff: "finite A \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 921 | setsum f A = Suc 0 \<longleftrightarrow> (EX a:A. f a = Suc 0 & (ALL b:A. a\<noteq>b \<longrightarrow> f b = 0))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 922 | apply(erule finite_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 923 | apply (auto simp add:add_is_1) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 924 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 925 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 926 | lemmas setsum_eq_1_iff = setsum_eq_Suc0_iff[simplified One_nat_def[symmetric]] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 927 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 928 | lemma setsum_Un_nat: "finite A ==> finite B ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 929 | (setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 930 |   -- {* For the natural numbers, we have subtraction. *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 931 | by (subst setsum_Un_Int [symmetric], auto simp add: algebra_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 932 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 933 | lemma setsum_diff1_nat: "(setsum f (A - {a}) :: nat) =
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 934 | (if a:A then setsum f A - f a else setsum f A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 935 | apply (case_tac "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 936 | prefer 2 apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 937 | apply (erule finite_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 938 | apply (auto simp add: insert_Diff_if) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 939 | apply (drule_tac a = a in mk_disjoint_insert, auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 940 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 941 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 942 | lemma setsum_diff_nat: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 943 | assumes "finite B" and "B \<subseteq> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 944 | shows "(setsum f (A - B) :: nat) = (setsum f A) - (setsum f B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 945 | using assms | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 946 | proof induct | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 947 |   show "setsum f (A - {}) = (setsum f A) - (setsum f {})" by simp
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 948 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 949 | fix F x assume finF: "finite F" and xnotinF: "x \<notin> F" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 950 | and xFinA: "insert x F \<subseteq> A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 951 | and IH: "F \<subseteq> A \<Longrightarrow> setsum f (A - F) = setsum f A - setsum f F" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 952 | from xnotinF xFinA have xinAF: "x \<in> (A - F)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 953 |   from xinAF have A: "setsum f ((A - F) - {x}) = setsum f (A - F) - f x"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 954 | by (simp add: setsum_diff1_nat) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 955 | from xFinA have "F \<subseteq> A" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 956 | with IH have "setsum f (A - F) = setsum f A - setsum f F" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 957 |   with A have B: "setsum f ((A - F) - {x}) = setsum f A - setsum f F - f x"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 958 | by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 959 |   from xnotinF have "A - insert x F = (A - F) - {x}" by auto
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 960 | with B have C: "setsum f (A - insert x F) = setsum f A - setsum f F - f x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 961 | by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 962 | from finF xnotinF have "setsum f (insert x F) = setsum f F + f x" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 963 | with C have "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 964 | by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 965 | thus "setsum f (A - insert x F) = setsum f A - setsum f (insert x F)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 966 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 967 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 968 | lemma setsum_comp_morphism: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 969 | assumes "h 0 = 0" and "\<And>x y. h (x + y) = h x + h y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 970 | shows "setsum (h \<circ> g) A = h (setsum g A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 971 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 972 | case False then show ?thesis by (simp add: assms) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 973 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 974 | case True then show ?thesis by (induct A) (simp_all add: assms) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 975 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 976 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 977 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 978 | subsubsection {* Cardinality as special case of @{const setsum} *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 979 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 980 | lemma card_eq_setsum: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 981 | "card A = setsum (\<lambda>x. 1) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 982 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 983 | have "plus \<circ> (\<lambda>_. Suc 0) = (\<lambda>_. Suc)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 984 | by (simp add: fun_eq_iff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 985 | then have "Finite_Set.fold (plus \<circ> (\<lambda>_. Suc 0)) = Finite_Set.fold (\<lambda>_. Suc)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 986 | by (rule arg_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 987 | then have "Finite_Set.fold (plus \<circ> (\<lambda>_. Suc 0)) 0 A = Finite_Set.fold (\<lambda>_. Suc) 0 A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 988 | by (blast intro: fun_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 989 | then show ?thesis by (simp add: card.eq_fold setsum.eq_fold) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 990 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 991 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 992 | lemma setsum_constant [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 993 | "(\<Sum>x \<in> A. y) = of_nat (card A) * y" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 994 | apply (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 995 | apply (erule finite_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 996 | apply (auto simp add: algebra_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 997 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 998 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 999 | lemma setsum_bounded: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1000 |   assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> (K::'a::{semiring_1, ordered_ab_semigroup_add})"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1001 | shows "setsum f A \<le> of_nat (card A) * K" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1002 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1003 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1004 | thus ?thesis using le setsum_mono[where K=A and g = "%x. K"] by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1005 | next | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1006 | case False thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1007 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1008 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1009 | lemma card_UN_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1010 | assumes "finite I" and "\<forall>i\<in>I. finite (A i)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1011 |     and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1012 | shows "card (UNION I A) = (\<Sum>i\<in>I. card(A i))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1013 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1014 | have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)" by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1015 | with assms show ?thesis by (simp add: card_eq_setsum setsum_UN_disjoint del: setsum_constant) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1016 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1017 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1018 | lemma card_Union_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1019 | "finite C ==> (ALL A:C. finite A) ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1020 |    (ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {})
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1021 | ==> card (Union C) = setsum card C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1022 | apply (frule card_UN_disjoint [of C id]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1023 | apply (simp_all add: SUP_def id_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1024 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1025 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1026 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1027 | subsubsection {* Cardinality of products *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1028 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1029 | lemma card_SigmaI [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1030 | "\<lbrakk> finite A; ALL a:A. finite (B a) \<rbrakk> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1031 | \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1032 | by(simp add: card_eq_setsum setsum_Sigma del:setsum_constant) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1033 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1034 | (* | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1035 | lemma SigmaI_insert: "y \<notin> A ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1036 |   (SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1037 | by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1038 | *) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1039 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1040 | lemma card_cartesian_product: "card (A <*> B) = card(A) * card(B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1041 | by (cases "finite A \<and> finite B") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1042 | (auto simp add: card_eq_0_iff dest: finite_cartesian_productD1 finite_cartesian_productD2) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1043 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1044 | lemma card_cartesian_product_singleton:  "card({x} <*> A) = card(A)"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1045 | by (simp add: card_cartesian_product) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1046 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1047 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1048 | subsection {* Generalized product over a set *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1049 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1050 | context comm_monoid_mult | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1051 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1052 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1053 | definition setprod :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1054 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1055 | "setprod = comm_monoid_set.F times 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1056 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1057 | sublocale setprod!: comm_monoid_set times 1 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1058 | where | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1059 | "comm_monoid_set.F times 1 = setprod" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1060 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1061 | show "comm_monoid_set times 1" .. | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1062 | then interpret setprod!: comm_monoid_set times 1 . | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1063 | from setprod_def show "comm_monoid_set.F times 1 = setprod" by rule | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1064 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1065 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1066 | abbreviation | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1067 |   Setprod ("\<Prod>_" [1000] 999) where
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1068 | "\<Prod>A \<equiv> setprod (\<lambda>x. x) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1069 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1070 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1071 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1072 | syntax | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1073 |   "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3PROD _:_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1074 | syntax (xsymbols) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1075 |   "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1076 | syntax (HTML output) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1077 |   "_setprod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(3\<Prod>_\<in>_. _)" [0, 51, 10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1078 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1079 | translations -- {* Beware of argument permutation! *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1080 | "PROD i:A. b" == "CONST setprod (%i. b) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1081 | "\<Prod>i\<in>A. b" == "CONST setprod (%i. b) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1082 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1083 | text{* Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1084 |  @{text"\<Prod>x|P. e"}. *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1085 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1086 | syntax | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1087 |   "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3PROD _ |/ _./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1088 | syntax (xsymbols) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1089 |   "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1090 | syntax (HTML output) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1091 |   "_qsetprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a" ("(3\<Prod>_ | (_)./ _)" [0,0,10] 10)
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1092 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1093 | translations | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1094 |   "PROD x|P. t" => "CONST setprod (%x. t) {x. P}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1095 |   "\<Prod>x|P. t" => "CONST setprod (%x. t) {x. P}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1096 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1097 | text {* TODO These are candidates for generalization *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1098 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1099 | context comm_monoid_mult | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1100 | begin | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1101 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1102 | lemma setprod_reindex_id: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1103 | "inj_on f B ==> setprod f B = setprod id (f ` B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1104 | by (auto simp add: setprod.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1105 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1106 | lemma setprod_reindex_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1107 | "inj_on f A ==> B = f ` A ==> g = h \<circ> f ==> setprod h B = setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1108 | by (frule setprod.reindex, simp) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1109 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1110 | lemma strong_setprod_reindex_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1111 | assumes i: "inj_on f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1112 | and B: "B = f ` A" and eq: "\<And>x. x \<in> A \<Longrightarrow> g x = (h \<circ> f) x" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1113 | shows "setprod h B = setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1114 | proof- | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1115 | have "setprod h B = setprod (h o f) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1116 | by (simp add: B setprod.reindex [OF i, of h]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1117 | then show ?thesis apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1118 | apply (rule setprod.cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1119 | apply simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1120 | by (simp add: eq) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1121 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1122 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1123 | lemma setprod_Union_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1124 |   assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A Int B = {}" 
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1125 | shows "setprod f (Union C) = setprod (setprod f) C" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1126 | using assms by (fact setprod.Union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1127 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1128 | text{*Here we can eliminate the finiteness assumptions, by cases.*}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1129 | lemma setprod_cartesian_product: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1130 | "(\<Prod>x\<in>A. (\<Prod>y\<in> B. f x y)) = (\<Prod>(x,y)\<in>(A <*> B). f x y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1131 | by (fact setprod.cartesian_product) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1132 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1133 | lemma setprod_Un2: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1134 | assumes "finite (A \<union> B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1135 | shows "setprod f (A \<union> B) = setprod f (A - B) * setprod f (B - A) * setprod f (A \<inter> B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1136 | proof - | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1137 | have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1138 | by auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1139 | with assms show ?thesis by simp (subst setprod.union_disjoint, auto)+ | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1140 | qed | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1141 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1142 | end | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1143 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1144 | text {* TODO These are legacy *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1145 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1146 | lemma setprod_empty: "setprod f {} = 1"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1147 | by (fact setprod.empty) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1148 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1149 | lemma setprod_insert: "[| finite A; a \<notin> A |] ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1150 | setprod f (insert a A) = f a * setprod f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1151 | by (fact setprod.insert) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1152 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1153 | lemma setprod_infinite: "~ finite A ==> setprod f A = 1" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1154 | by (fact setprod.infinite) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1155 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1156 | lemma setprod_reindex: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1157 | "inj_on f B ==> setprod h (f ` B) = setprod (h \<circ> f) B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1158 | by (fact setprod.reindex) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1159 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1160 | lemma setprod_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1161 | "A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1162 | by (fact setprod.cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1163 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1164 | lemma strong_setprod_cong: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1165 | "A = B ==> (!!x. x:B =simp=> f x = g x) ==> setprod f A = setprod g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1166 | by (fact setprod.strong_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1167 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1168 | lemma setprod_Un_one: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1169 | "\<lbrakk> finite S; finite T; \<forall>x \<in> S\<inter>T. f x = 1 \<rbrakk> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1170 | \<Longrightarrow> setprod f (S \<union> T) = setprod f S * setprod f T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1171 | by (fact setprod.union_inter_neutral) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1172 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1173 | lemmas setprod_1 = setprod.neutral_const | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1174 | lemmas setprod_1' = setprod.neutral | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1175 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1176 | lemma setprod_Un_Int: "finite A ==> finite B | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1177 | ==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1178 | by (fact setprod.union_inter) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1179 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1180 | lemma setprod_Un_disjoint: "finite A ==> finite B | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1181 |   ==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1182 | by (fact setprod.union_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1183 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1184 | lemma setprod_subset_diff: "\<lbrakk> B \<subseteq> A; finite A \<rbrakk> \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1185 | setprod f A = setprod f (A - B) * setprod f B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1186 | by (fact setprod.subset_diff) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1187 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1188 | lemma setprod_mono_one_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1189 | "\<lbrakk> finite T; S \<subseteq> T; \<forall>i \<in> T - S. f i = 1 \<rbrakk> \<Longrightarrow> setprod f S = setprod f T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1190 | by (fact setprod.mono_neutral_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1191 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1192 | lemmas setprod_mono_one_right = setprod.mono_neutral_right | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1193 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1194 | lemma setprod_mono_one_cong_left: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1195 | "\<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> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1196 | \<Longrightarrow> setprod f S = setprod g T" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1197 | by (fact setprod.mono_neutral_cong_left) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1198 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1199 | lemmas setprod_mono_one_cong_right = setprod.mono_neutral_cong_right | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1200 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1201 | lemma setprod_delta: "finite S \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1202 | setprod (\<lambda>k. if k=a then b k else 1) S = (if a \<in> S then b a else 1)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1203 | by (fact setprod.delta) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1204 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1205 | lemma setprod_delta': "finite S \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1206 | setprod (\<lambda>k. if a = k then b k else 1) S = (if a\<in> S then b a else 1)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1207 | by (fact setprod.delta') | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1208 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1209 | lemma setprod_UN_disjoint: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1210 | "finite I ==> (ALL i:I. finite (A i)) ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1211 |         (ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1212 | setprod f (UNION I A) = setprod (%i. setprod f (A i)) I" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1213 | by (fact setprod.UNION_disjoint) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1214 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1215 | lemma setprod_Sigma: "finite A ==> ALL x:A. finite (B x) ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1216 | (\<Prod>x\<in>A. (\<Prod>y\<in> B x. f x y)) = | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1217 | (\<Prod>(x,y)\<in>(SIGMA x:A. B x). f x y)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1218 | by (fact setprod.Sigma) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1219 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1220 | lemma setprod_timesf: "setprod (\<lambda>x. f x * g x) A = setprod f A * setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1221 | by (fact setprod.distrib) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1222 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1223 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1224 | subsubsection {* Properties in more restricted classes of structures *}
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1225 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1226 | lemma setprod_zero: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1227 | "finite A ==> EX x: A. f x = (0::'a::comm_semiring_1) ==> setprod f A = 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1228 | apply (induct set: finite, force, clarsimp) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1229 | apply (erule disjE, auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1230 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1231 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1232 | lemma setprod_zero_iff[simp]: "finite A ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1233 |   (setprod f A = (0::'a::{comm_semiring_1,no_zero_divisors})) =
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1234 | (EX x: A. f x = 0)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1235 | by (erule finite_induct, auto simp:no_zero_divisors) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1236 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1237 | lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1238 |   (setprod f (A Un B) :: 'a ::{field})
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1239 | = setprod f A * setprod f B / setprod f (A Int B)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1240 | by (subst setprod_Un_Int [symmetric], auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1241 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1242 | lemma setprod_nonneg [rule_format]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1243 | "(ALL x: A. (0::'a::linordered_semidom) \<le> f x) --> 0 \<le> setprod f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1244 | by (cases "finite A", induct set: finite, simp_all add: mult_nonneg_nonneg) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1245 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1246 | lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::linordered_semidom) < f x) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1247 | --> 0 < setprod f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1248 | by (cases "finite A", induct set: finite, simp_all add: mult_pos_pos) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1249 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1250 | lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1251 |   (setprod f (A - {a}) :: 'a :: {field}) =
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1252 | (if a:A then setprod f A / f a else setprod f A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1253 | by (erule finite_induct) (auto simp add: insert_Diff_if) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1254 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1255 | lemma setprod_inversef: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1256 | fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1257 | shows "finite A ==> setprod (inverse \<circ> f) A = inverse (setprod f A)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1258 | by (erule finite_induct) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1259 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1260 | lemma setprod_dividef: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1261 | fixes f :: "'b \<Rightarrow> 'a::field_inverse_zero" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1262 | shows "finite A | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1263 | ==> setprod (%x. f x / g x) A = setprod f A / setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1264 | apply (subgoal_tac | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1265 | "setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1266 | apply (erule ssubst) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1267 | apply (subst divide_inverse) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1268 | apply (subst setprod_timesf) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1269 | apply (subst setprod_inversef, assumption+, rule refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1270 | apply (rule setprod_cong, rule refl) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1271 | apply (subst divide_inverse, auto) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1272 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1273 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1274 | lemma setprod_dvd_setprod [rule_format]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1275 | "(ALL x : A. f x dvd g x) \<longrightarrow> setprod f A dvd setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1276 | apply (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1277 | apply (induct set: finite) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1278 | apply (auto simp add: dvd_def) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1279 | apply (rule_tac x = "k * ka" in exI) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1280 | apply (simp add: algebra_simps) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1281 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1282 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1283 | lemma setprod_dvd_setprod_subset: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1284 | "finite B \<Longrightarrow> A <= B \<Longrightarrow> setprod f A dvd setprod f B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1285 | apply (subgoal_tac "setprod f B = setprod f A * setprod f (B - A)") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1286 | apply (unfold dvd_def, blast) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1287 | apply (subst setprod_Un_disjoint [symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1288 | apply (auto elim: finite_subset intro: setprod_cong) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1289 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1290 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1291 | lemma setprod_dvd_setprod_subset2: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1292 | "finite B \<Longrightarrow> A <= B \<Longrightarrow> ALL x : A. (f x::'a::comm_semiring_1) dvd g x \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1293 | setprod f A dvd setprod g B" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1294 | apply (rule dvd_trans) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1295 | apply (rule setprod_dvd_setprod, erule (1) bspec) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1296 | apply (erule (1) setprod_dvd_setprod_subset) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1297 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1298 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1299 | lemma dvd_setprod: "finite A \<Longrightarrow> i:A \<Longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1300 | (f i ::'a::comm_semiring_1) dvd setprod f A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1301 | by (induct set: finite) (auto intro: dvd_mult) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1302 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1303 | lemma dvd_setsum [rule_format]: "(ALL i : A. d dvd f i) \<longrightarrow> | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1304 | (d::'a::comm_semiring_1) dvd (SUM x : A. f x)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1305 | apply (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1306 | apply (induct set: finite) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1307 | apply auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1308 | done | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1309 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1310 | lemma setprod_mono: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1311 | fixes f :: "'a \<Rightarrow> 'b\<Colon>linordered_semidom" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1312 | assumes "\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1313 | shows "setprod f A \<le> setprod g A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1314 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1315 | case True | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1316 | hence ?thesis "setprod f A \<ge> 0" using subset_refl[of A] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1317 | proof (induct A rule: finite_subset_induct) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1318 | case (insert a F) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1319 | thus "setprod f (insert a F) \<le> setprod g (insert a F)" "0 \<le> setprod f (insert a F)" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1320 | unfolding setprod_insert[OF insert(1,3)] | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1321 | using assms[rule_format,OF insert(2)] insert | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1322 | by (auto intro: mult_mono mult_nonneg_nonneg) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1323 | qed auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1324 | thus ?thesis by simp | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1325 | qed auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1326 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1327 | lemma abs_setprod: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1328 |   fixes f :: "'a \<Rightarrow> 'b\<Colon>{linordered_field,abs}"
 | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1329 | shows "abs (setprod f A) = setprod (\<lambda>x. abs (f x)) A" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1330 | proof (cases "finite A") | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1331 | case True thus ?thesis | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1332 | by induct (auto simp add: field_simps abs_mult) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1333 | qed auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1334 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1335 | lemma setprod_eq_1_iff [simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1336 | "finite F ==> setprod f F = 1 \<longleftrightarrow> (ALL a:F. f a = (1::nat))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1337 | by (induct set: finite) auto | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1338 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1339 | lemma setprod_pos_nat: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1340 | "finite S ==> (ALL x : S. f x > (0::nat)) ==> setprod f S > 0" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1341 | using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1342 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1343 | lemma setprod_pos_nat_iff[simp]: | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1344 | "finite S ==> (setprod f S > 0) = (ALL x : S. f x > (0::nat))" | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1345 | using setprod_zero_iff by(simp del:neq0_conv add:neq0_conv[symmetric]) | 
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1346 | |
| 
1e7f2d296e19
more algebraic terminology for theories about big operators
 haftmann parents: diff
changeset | 1347 | end |