| author | blanchet | 
| Tue, 02 Jan 2018 16:01:03 +0100 | |
| changeset 67316 | adaf279ce67b | 
| parent 67268 | bdf25939a550 | 
| child 67511 | a6f5a78712af | 
| permissions | -rw-r--r-- | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Groups_Big.thy  | 
| 63654 | 2  | 
Author: Tobias Nipkow  | 
3  | 
Author: Lawrence C Paulson  | 
|
4  | 
Author: Markus Wenzel  | 
|
5  | 
Author: Jeremy Avigad  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
7  | 
|
| 60758 | 8  | 
section \<open>Big sum and product over finite (non-empty) sets\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
10  | 
theory Groups_Big  | 
| 66364 | 11  | 
imports Power  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
13  | 
|
| 60758 | 14  | 
subsection \<open>Generic monoid operation over a set\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
16  | 
locale comm_monoid_set = comm_monoid  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
17  | 
begin  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
19  | 
interpretation comp_fun_commute f  | 
| 61169 | 20  | 
by standard (simp add: fun_eq_iff left_commute)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
21  | 
|
| 54745 | 22  | 
interpretation comp?: comp_fun_commute "f \<circ> g"  | 
23  | 
by (fact comp_comp_fun_commute)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
25  | 
definition F :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
 | 
| 63654 | 26  | 
where eq_fold: "F g A = Finite_Set.fold (f \<circ> g) \<^bold>1 A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
27  | 
|
| 63654 | 28  | 
lemma infinite [simp]: "\<not> finite A \<Longrightarrow> F g A = \<^bold>1"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
29  | 
by (simp add: eq_fold)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
30  | 
|
| 63654 | 31  | 
lemma empty [simp]: "F g {} = \<^bold>1"
 | 
| 
54744
 
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  | 
|
| 63654 | 34  | 
lemma insert [simp]: "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> F g (insert x A) = g x \<^bold>* F g A"  | 
35  | 
by (simp add: eq_fold)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
37  | 
lemma remove:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
38  | 
assumes "finite A" and "x \<in> A"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
39  | 
  shows "F g A = g x \<^bold>* F g (A - {x})"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
40  | 
proof -  | 
| 63654 | 41  | 
from \<open>x \<in> A\<close> obtain B where B: "A = insert x B" and "x \<notin> B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
42  | 
by (auto dest: mk_disjoint_insert)  | 
| 63654 | 43  | 
moreover from \<open>finite A\<close> B have "finite B" by simp  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
44  | 
ultimately show ?thesis by simp  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
45  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
46  | 
|
| 63654 | 47  | 
lemma insert_remove: "finite A \<Longrightarrow> F g (insert x A) = g x \<^bold>* F g (A - {x})"
 | 
48  | 
by (cases "x \<in> A") (simp_all add: remove insert_absorb)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
49  | 
|
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
50  | 
lemma insert_if: "finite A \<Longrightarrow> F g (insert x A) = (if x \<in> A then F g A else g x \<^bold>* F g A)"  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
51  | 
by (cases "x \<in> A") (simp_all add: insert_absorb)  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
52  | 
|
| 63654 | 53  | 
lemma neutral: "\<forall>x\<in>A. g x = \<^bold>1 \<Longrightarrow> F g A = \<^bold>1"  | 
54  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
55  | 
|
| 63654 | 56  | 
lemma neutral_const [simp]: "F (\<lambda>_. \<^bold>1) A = \<^bold>1"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
57  | 
by (simp add: neutral)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
59  | 
lemma union_inter:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
60  | 
assumes "finite A" and "finite B"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
61  | 
shows "F g (A \<union> B) \<^bold>* F g (A \<inter> B) = F g A \<^bold>* F g B"  | 
| 61799 | 62  | 
\<comment> \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close>  | 
| 63654 | 63  | 
using assms  | 
64  | 
proof (induct A)  | 
|
65  | 
case empty  | 
|
66  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
67  | 
next  | 
| 63654 | 68  | 
case (insert x A)  | 
69  | 
then show ?case  | 
|
70  | 
by (auto simp: insert_absorb Int_insert_left commute [of _ "g x"] assoc left_commute)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
71  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
73  | 
corollary union_inter_neutral:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
74  | 
assumes "finite A" and "finite B"  | 
| 63654 | 75  | 
and "\<forall>x \<in> A \<inter> B. g x = \<^bold>1"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
76  | 
shows "F g (A \<union> B) = F g A \<^bold>* F g B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
77  | 
using assms by (simp add: union_inter [symmetric] neutral)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
79  | 
corollary union_disjoint:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
80  | 
assumes "finite A" and "finite B"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
81  | 
  assumes "A \<inter> B = {}"
 | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
82  | 
shows "F g (A \<union> B) = F g A \<^bold>* F g B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
83  | 
using assms by (simp add: union_inter_neutral)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
84  | 
|
| 57418 | 85  | 
lemma union_diff2:  | 
86  | 
assumes "finite A" and "finite B"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
87  | 
shows "F g (A \<union> B) = F g (A - B) \<^bold>* F g (B - A) \<^bold>* F g (A \<inter> B)"  | 
| 57418 | 88  | 
proof -  | 
89  | 
have "A \<union> B = A - B \<union> (B - A) \<union> A \<inter> B"  | 
|
90  | 
by auto  | 
|
| 63654 | 91  | 
with assms show ?thesis  | 
92  | 
by simp (subst union_disjoint, auto)+  | 
|
| 57418 | 93  | 
qed  | 
94  | 
||
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
95  | 
lemma subset_diff:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
96  | 
assumes "B \<subseteq> A" and "finite A"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
97  | 
shows "F g A = F g (A - B) \<^bold>* F g B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
98  | 
proof -  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
99  | 
from assms have "finite (A - B)" by auto  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
100  | 
moreover from assms have "finite B" by (rule finite_subset)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
101  | 
  moreover from assms have "(A - B) \<inter> B = {}" by auto
 | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
102  | 
ultimately have "F g (A - B \<union> B) = F g (A - B) \<^bold>* F g B" by (rule union_disjoint)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
103  | 
moreover from assms have "A \<union> B = A" by auto  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
104  | 
ultimately show ?thesis by simp  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
105  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
106  | 
|
| 56545 | 107  | 
lemma setdiff_irrelevant:  | 
108  | 
assumes "finite A"  | 
|
109  | 
  shows "F g (A - {x. g x = z}) = F g A"
 | 
|
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
110  | 
using assms by (induct A) (simp_all add: insert_Diff_if)  | 
| 58195 | 111  | 
|
| 56545 | 112  | 
lemma not_neutral_contains_not_neutral:  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
113  | 
assumes "F g A \<noteq> \<^bold>1"  | 
| 
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
114  | 
obtains a where "a \<in> A" and "g a \<noteq> \<^bold>1"  | 
| 56545 | 115  | 
proof -  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
116  | 
from assms have "\<exists>a\<in>A. g a \<noteq> \<^bold>1"  | 
| 56545 | 117  | 
proof (induct A rule: infinite_finite_induct)  | 
| 63654 | 118  | 
case infinite  | 
119  | 
then show ?case by simp  | 
|
120  | 
next  | 
|
121  | 
case empty  | 
|
122  | 
then show ?case by simp  | 
|
123  | 
next  | 
|
| 56545 | 124  | 
case (insert a A)  | 
| 63654 | 125  | 
then show ?case by fastforce  | 
126  | 
qed  | 
|
| 56545 | 127  | 
with that show thesis by blast  | 
128  | 
qed  | 
|
129  | 
||
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
130  | 
lemma reindex:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
131  | 
assumes "inj_on h A"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
132  | 
shows "F g (h ` A) = F (g \<circ> h) A"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
133  | 
proof (cases "finite A")  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
134  | 
case True  | 
| 63654 | 135  | 
with assms show ?thesis  | 
136  | 
by (simp add: eq_fold fold_image comp_assoc)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
137  | 
next  | 
| 63654 | 138  | 
case False  | 
139  | 
with assms have "\<not> finite (h ` A)" by (blast dest: finite_imageD)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
140  | 
with False show ?thesis by simp  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
141  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
142  | 
|
| 63357 | 143  | 
lemma cong [fundef_cong]:  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
144  | 
assumes "A = B"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
145  | 
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
 | 
146  | 
shows "F g A = F h B"  | 
| 60758 | 147  | 
using g_h unfolding \<open>A = B\<close>  | 
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
148  | 
by (induct B rule: infinite_finite_induct) auto  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
150  | 
lemma strong_cong [cong]:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
151  | 
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
 | 
152  | 
shows "F (\<lambda>x. g x) A = F (\<lambda>x. h x) B"  | 
| 63654 | 153  | 
by (rule cong) (use assms in \<open>simp_all add: simp_implies_def\<close>)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
154  | 
|
| 57418 | 155  | 
lemma reindex_cong:  | 
156  | 
assumes "inj_on l B"  | 
|
157  | 
assumes "A = l ` B"  | 
|
158  | 
assumes "\<And>x. x \<in> B \<Longrightarrow> g (l x) = h x"  | 
|
159  | 
shows "F g A = F h B"  | 
|
160  | 
using assms by (simp add: reindex)  | 
|
161  | 
||
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
162  | 
lemma UNION_disjoint:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
163  | 
assumes "finite I" and "\<forall>i\<in>I. finite (A i)"  | 
| 63654 | 164  | 
    and "\<forall>i\<in>I. \<forall>j\<in>I. i \<noteq> j \<longrightarrow> A i \<inter> A j = {}"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
165  | 
shows "F g (UNION I A) = F (\<lambda>x. F g (A x)) I"  | 
| 63654 | 166  | 
apply (insert assms)  | 
167  | 
apply (induct rule: finite_induct)  | 
|
168  | 
apply simp  | 
|
169  | 
apply atomize  | 
|
170  | 
apply (subgoal_tac "\<forall>i\<in>Fa. x \<noteq> i")  | 
|
171  | 
prefer 2 apply blast  | 
|
172  | 
  apply (subgoal_tac "A x \<inter> UNION Fa A = {}")
 | 
|
173  | 
prefer 2 apply blast  | 
|
174  | 
apply (simp add: union_disjoint)  | 
|
175  | 
done  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
176  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
177  | 
lemma Union_disjoint:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
178  | 
  assumes "\<forall>A\<in>C. finite A" "\<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {}"
 | 
| 61952 | 179  | 
shows "F g (\<Union>C) = (F \<circ> F) g C"  | 
| 63654 | 180  | 
proof (cases "finite C")  | 
181  | 
case True  | 
|
182  | 
from UNION_disjoint [OF this assms] show ?thesis by simp  | 
|
183  | 
next  | 
|
184  | 
case False  | 
|
185  | 
then show ?thesis by (auto dest: finite_UnionD intro: infinite)  | 
|
186  | 
qed  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
187  | 
|
| 63654 | 188  | 
lemma distrib: "F (\<lambda>x. g x \<^bold>* h x) A = F g A \<^bold>* F h A"  | 
| 63092 | 189  | 
by (induct A rule: infinite_finite_induct) (simp_all add: assoc commute left_commute)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
191  | 
lemma Sigma:  | 
| 
61032
 
b57df8eecad6
standardized some occurences of ancient "split" alias
 
haftmann 
parents: 
60974 
diff
changeset
 | 
192  | 
"finite A \<Longrightarrow> \<forall>x\<in>A. finite (B x) \<Longrightarrow> F (\<lambda>x. F (g x) (B x)) A = F (case_prod g) (SIGMA x:A. B x)"  | 
| 63654 | 193  | 
apply (subst Sigma_def)  | 
194  | 
apply (subst UNION_disjoint)  | 
|
195  | 
apply assumption  | 
|
196  | 
apply simp  | 
|
197  | 
apply blast  | 
|
198  | 
apply (rule cong)  | 
|
199  | 
apply rule  | 
|
200  | 
apply (simp add: fun_eq_iff)  | 
|
201  | 
apply (subst UNION_disjoint)  | 
|
202  | 
apply simp  | 
|
203  | 
apply simp  | 
|
204  | 
apply blast  | 
|
205  | 
apply (simp add: comp_def)  | 
|
206  | 
done  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
207  | 
|
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
208  | 
lemma related:  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
209  | 
assumes Re: "R \<^bold>1 \<^bold>1"  | 
| 63654 | 210  | 
and Rop: "\<forall>x1 y1 x2 y2. R x1 x2 \<and> R y1 y2 \<longrightarrow> R (x1 \<^bold>* y1) (x2 \<^bold>* y2)"  | 
211  | 
and fin: "finite S"  | 
|
212  | 
and R_h_g: "\<forall>x\<in>S. R (h x) (g x)"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
213  | 
shows "R (F h S) (F g S)"  | 
| 63654 | 214  | 
using fin by (rule finite_subset_induct) (use assms in auto)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
216  | 
lemma mono_neutral_cong_left:  | 
| 63654 | 217  | 
assumes "finite T"  | 
218  | 
and "S \<subseteq> T"  | 
|
219  | 
and "\<forall>i \<in> T - S. h i = \<^bold>1"  | 
|
220  | 
and "\<And>x. x \<in> S \<Longrightarrow> g x = h x"  | 
|
221  | 
shows "F g S = F h T"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
222  | 
proof-  | 
| 60758 | 223  | 
have eq: "T = S \<union> (T - S)" using \<open>S \<subseteq> T\<close> by blast  | 
224  | 
  have d: "S \<inter> (T - S) = {}" using \<open>S \<subseteq> T\<close> by blast
 | 
|
225  | 
from \<open>finite T\<close> \<open>S \<subseteq> T\<close> have f: "finite S" "finite (T - S)"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
226  | 
by (auto intro: finite_subset)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
227  | 
show ?thesis using assms(4)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
228  | 
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
 | 
229  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
231  | 
lemma mono_neutral_cong_right:  | 
| 63654 | 232  | 
"finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. g i = \<^bold>1 \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> g x = h x) \<Longrightarrow>  | 
233  | 
F g T = F h S"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
234  | 
by (auto intro!: mono_neutral_cong_left [symmetric])  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
235  | 
|
| 63654 | 236  | 
lemma mono_neutral_left: "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. g i = \<^bold>1 \<Longrightarrow> F g S = F g T"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
237  | 
by (blast intro: mono_neutral_cong_left)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
238  | 
|
| 63654 | 239  | 
lemma mono_neutral_right: "finite T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> \<forall>i \<in> T - S. g i = \<^bold>1 \<Longrightarrow> F g T = F g S"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
240  | 
by (blast intro!: mono_neutral_left [symmetric])  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
241  | 
|
| 
64979
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
242  | 
lemma mono_neutral_cong:  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
243  | 
assumes [simp]: "finite T" "finite S"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
244  | 
and *: "\<And>i. i \<in> T - S \<Longrightarrow> h i = \<^bold>1" "\<And>i. i \<in> S - T \<Longrightarrow> g i = \<^bold>1"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
245  | 
and gh: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x = h x"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
246  | 
shows "F g S = F h T"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
247  | 
proof-  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
248  | 
have "F g S = F g (S \<inter> T)"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
249  | 
by(rule mono_neutral_right)(auto intro: *)  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
250  | 
also have "\<dots> = F h (S \<inter> T)" using refl gh by(rule cong)  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
251  | 
also have "\<dots> = F h T"  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
252  | 
by(rule mono_neutral_left)(auto intro: *)  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
253  | 
finally show ?thesis .  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
254  | 
qed  | 
| 
 
20a623d03d71
move mono_neutral_cong from AFP/Deep_Learning/PP_Auxiliary
 
hoelzl 
parents: 
64272 
diff
changeset
 | 
255  | 
|
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
256  | 
lemma reindex_bij_betw: "bij_betw h S T \<Longrightarrow> F (\<lambda>x. g (h x)) S = F g T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
257  | 
by (auto simp: bij_betw_def reindex)  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
258  | 
|
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
259  | 
lemma reindex_bij_witness:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
260  | 
assumes witness:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
261  | 
"\<And>a. a \<in> S \<Longrightarrow> i (j a) = a"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
262  | 
"\<And>a. a \<in> S \<Longrightarrow> j a \<in> T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
263  | 
"\<And>b. b \<in> T \<Longrightarrow> j (i b) = b"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
264  | 
"\<And>b. b \<in> T \<Longrightarrow> i b \<in> S"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
265  | 
assumes eq:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
266  | 
"\<And>a. a \<in> S \<Longrightarrow> h (j a) = g a"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
267  | 
shows "F g S = F h T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
268  | 
proof -  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
269  | 
have "bij_betw j S T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
270  | 
using bij_betw_byWitness[where A=S and f=j and f'=i and A'=T] witness by auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
271  | 
moreover have "F g S = F (\<lambda>x. h (j x)) S"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
272  | 
by (intro cong) (auto simp: eq)  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
273  | 
ultimately show ?thesis  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
274  | 
by (simp add: reindex_bij_betw)  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
275  | 
qed  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
276  | 
|
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
277  | 
lemma reindex_bij_betw_not_neutral:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
278  | 
assumes fin: "finite S'" "finite T'"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
279  | 
assumes bij: "bij_betw h (S - S') (T - T')"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
280  | 
assumes nn:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
281  | 
"\<And>a. a \<in> S' \<Longrightarrow> g (h a) = z"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
282  | 
"\<And>b. b \<in> T' \<Longrightarrow> g b = z"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
283  | 
shows "F (\<lambda>x. g (h x)) S = F g T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
284  | 
proof -  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
285  | 
have [simp]: "finite S \<longleftrightarrow> finite T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
286  | 
using bij_betw_finite[OF bij] fin by auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
287  | 
show ?thesis  | 
| 63654 | 288  | 
proof (cases "finite S")  | 
289  | 
case True  | 
|
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
290  | 
with nn have "F (\<lambda>x. g (h x)) S = F (\<lambda>x. g (h x)) (S - S')"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
291  | 
by (intro mono_neutral_cong_right) auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
292  | 
also have "\<dots> = F g (T - T')"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
293  | 
using bij by (rule reindex_bij_betw)  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
294  | 
also have "\<dots> = F g T"  | 
| 60758 | 295  | 
using nn \<open>finite S\<close> by (intro mono_neutral_cong_left) auto  | 
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
296  | 
finally show ?thesis .  | 
| 63654 | 297  | 
next  | 
298  | 
case False  | 
|
299  | 
then show ?thesis by simp  | 
|
300  | 
qed  | 
|
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
301  | 
qed  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
302  | 
|
| 57418 | 303  | 
lemma reindex_nontrivial:  | 
304  | 
assumes "finite A"  | 
|
| 63654 | 305  | 
and nz: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<noteq> y \<Longrightarrow> h x = h y \<Longrightarrow> g (h x) = \<^bold>1"  | 
| 57418 | 306  | 
shows "F g (h ` A) = F (g \<circ> h) A"  | 
307  | 
proof (subst reindex_bij_betw_not_neutral [symmetric])  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
308  | 
  show "bij_betw h (A - {x \<in> A. (g \<circ> h) x = \<^bold>1}) (h ` A - h ` {x \<in> A. (g \<circ> h) x = \<^bold>1})"
 | 
| 57418 | 309  | 
using nz by (auto intro!: inj_onI simp: bij_betw_def)  | 
| 63654 | 310  | 
qed (use \<open>finite A\<close> in auto)  | 
| 57418 | 311  | 
|
| 
57129
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
312  | 
lemma reindex_bij_witness_not_neutral:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
313  | 
assumes fin: "finite S'" "finite T'"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
314  | 
assumes witness:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
315  | 
"\<And>a. a \<in> S - S' \<Longrightarrow> i (j a) = a"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
316  | 
"\<And>a. a \<in> S - S' \<Longrightarrow> j a \<in> T - T'"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
317  | 
"\<And>b. b \<in> T - T' \<Longrightarrow> j (i b) = b"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
318  | 
"\<And>b. b \<in> T - T' \<Longrightarrow> i b \<in> S - S'"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
319  | 
assumes nn:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
320  | 
"\<And>a. a \<in> S' \<Longrightarrow> g a = z"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
321  | 
"\<And>b. b \<in> T' \<Longrightarrow> h b = z"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
322  | 
assumes eq:  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
323  | 
"\<And>a. a \<in> S \<Longrightarrow> h (j a) = g a"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
324  | 
shows "F g S = F h T"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
325  | 
proof -  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
326  | 
have bij: "bij_betw j (S - (S' \<inter> S)) (T - (T' \<inter> T))"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
327  | 
using witness by (intro bij_betw_byWitness[where f'=i]) auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
328  | 
have F_eq: "F g S = F (\<lambda>x. h (j x)) S"  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
329  | 
by (intro cong) (auto simp: eq)  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
330  | 
show ?thesis  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
331  | 
unfolding F_eq using fin nn eq  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
332  | 
by (intro reindex_bij_betw_not_neutral[OF _ _ bij]) auto  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
333  | 
qed  | 
| 
 
7edb7550663e
introduce more powerful reindexing rules for big operators
 
hoelzl 
parents: 
56545 
diff
changeset
 | 
334  | 
|
| 
66089
 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 
paulson <lp15@cam.ac.uk> 
parents: 
65687 
diff
changeset
 | 
335  | 
lemma delta [simp]:  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
336  | 
assumes fS: "finite S"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
337  | 
shows "F (\<lambda>k. if k = a then b k else \<^bold>1) S = (if a \<in> S then b a else \<^bold>1)"  | 
| 63654 | 338  | 
proof -  | 
339  | 
let ?f = "(\<lambda>k. if k = a then b k else \<^bold>1)"  | 
|
340  | 
show ?thesis  | 
|
341  | 
proof (cases "a \<in> S")  | 
|
342  | 
case False  | 
|
343  | 
then have "\<forall>k\<in>S. ?f k = \<^bold>1" by simp  | 
|
344  | 
with False show ?thesis by simp  | 
|
345  | 
next  | 
|
346  | 
case True  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
347  | 
    let ?A = "S - {a}"
 | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
348  | 
    let ?B = "{a}"
 | 
| 63654 | 349  | 
from True have eq: "S = ?A \<union> ?B" by blast  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
350  | 
    have dj: "?A \<inter> ?B = {}" by simp
 | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
351  | 
from fS have fAB: "finite ?A" "finite ?B" by auto  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
352  | 
have "F ?f S = F ?f ?A \<^bold>* F ?f ?B"  | 
| 63654 | 353  | 
using union_disjoint [OF fAB dj, of ?f, unfolded eq [symmetric]] by simp  | 
354  | 
with True show ?thesis by simp  | 
|
355  | 
qed  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
356  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
357  | 
|
| 
66089
 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 
paulson <lp15@cam.ac.uk> 
parents: 
65687 
diff
changeset
 | 
358  | 
lemma delta' [simp]:  | 
| 63654 | 359  | 
assumes fin: "finite S"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
360  | 
shows "F (\<lambda>k. if a = k then b k else \<^bold>1) S = (if a \<in> S then b a else \<^bold>1)"  | 
| 63654 | 361  | 
using delta [OF fin, of a b, symmetric] by (auto intro: cong)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
362  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
363  | 
lemma If_cases:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
364  | 
fixes P :: "'b \<Rightarrow> bool" and g h :: "'b \<Rightarrow> 'a"  | 
| 63654 | 365  | 
assumes fin: "finite A"  | 
366  | 
  shows "F (\<lambda>x. if P x then h x else g x) A = F h (A \<inter> {x. P x}) \<^bold>* F g (A \<inter> - {x. P x})"
 | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
367  | 
proof -  | 
| 63654 | 368  | 
  have a: "A = A \<inter> {x. P x} \<union> A \<inter> -{x. P x}" "(A \<inter> {x. P x}) \<inter> (A \<inter> -{x. P x}) = {}"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
369  | 
by blast+  | 
| 63654 | 370  | 
  from fin have f: "finite (A \<inter> {x. P x})" "finite (A \<inter> -{x. P x})" by auto
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
371  | 
let ?g = "\<lambda>x. if P x then h x else g x"  | 
| 63654 | 372  | 
from union_disjoint [OF f a(2), of ?g] a(1) show ?thesis  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
373  | 
by (subst (1 2) cong) simp_all  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
374  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
375  | 
|
| 63654 | 376  | 
lemma cartesian_product: "F (\<lambda>x. F (g x) B) A = F (case_prod g) (A \<times> B)"  | 
377  | 
apply (rule sym)  | 
|
378  | 
apply (cases "finite A")  | 
|
379  | 
apply (cases "finite B")  | 
|
380  | 
apply (simp add: Sigma)  | 
|
381  | 
   apply (cases "A = {}")
 | 
|
382  | 
apply simp  | 
|
383  | 
apply simp  | 
|
384  | 
apply (auto intro: infinite dest: finite_cartesian_productD2)  | 
|
385  | 
  apply (cases "B = {}")
 | 
|
386  | 
apply (auto intro: infinite dest: finite_cartesian_productD1)  | 
|
387  | 
done  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
388  | 
|
| 57418 | 389  | 
lemma inter_restrict:  | 
390  | 
assumes "finite A"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
391  | 
shows "F g (A \<inter> B) = F (\<lambda>x. if x \<in> B then g x else \<^bold>1) A"  | 
| 57418 | 392  | 
proof -  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
393  | 
let ?g = "\<lambda>x. if x \<in> A \<inter> B then g x else \<^bold>1"  | 
| 63654 | 394  | 
have "\<forall>i\<in>A - A \<inter> B. (if i \<in> A \<inter> B then g i else \<^bold>1) = \<^bold>1" by simp  | 
| 57418 | 395  | 
moreover have "A \<inter> B \<subseteq> A" by blast  | 
| 63654 | 396  | 
ultimately have "F ?g (A \<inter> B) = F ?g A"  | 
397  | 
using \<open>finite A\<close> by (intro mono_neutral_left) auto  | 
|
| 57418 | 398  | 
then show ?thesis by simp  | 
399  | 
qed  | 
|
400  | 
||
401  | 
lemma inter_filter:  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
402  | 
  "finite A \<Longrightarrow> F g {x \<in> A. P x} = F (\<lambda>x. if P x then g x else \<^bold>1) A"
 | 
| 57418 | 403  | 
  by (simp add: inter_restrict [symmetric, of A "{x. P x}" g, simplified mem_Collect_eq] Int_def)
 | 
404  | 
||
405  | 
lemma Union_comp:  | 
|
406  | 
assumes "\<forall>A \<in> B. finite A"  | 
|
| 63654 | 407  | 
and "\<And>A1 A2 x. A1 \<in> B \<Longrightarrow> A2 \<in> B \<Longrightarrow> A1 \<noteq> A2 \<Longrightarrow> x \<in> A1 \<Longrightarrow> x \<in> A2 \<Longrightarrow> g x = \<^bold>1"  | 
| 57418 | 408  | 
shows "F g (\<Union>B) = (F \<circ> F) g B"  | 
| 63654 | 409  | 
using assms  | 
410  | 
proof (induct B rule: infinite_finite_induct)  | 
|
| 57418 | 411  | 
case (infinite A)  | 
412  | 
then have "\<not> finite (\<Union>A)" by (blast dest: finite_UnionD)  | 
|
413  | 
with infinite show ?case by simp  | 
|
414  | 
next  | 
|
| 63654 | 415  | 
case empty  | 
416  | 
then show ?case by simp  | 
|
| 57418 | 417  | 
next  | 
418  | 
case (insert A B)  | 
|
419  | 
then have "finite A" "finite B" "finite (\<Union>B)" "A \<notin> B"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
420  | 
and "\<forall>x\<in>A \<inter> \<Union>B. g x = \<^bold>1"  | 
| 63654 | 421  | 
and H: "F g (\<Union>B) = (F \<circ> F) g B" by auto  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
422  | 
then have "F g (A \<union> \<Union>B) = F g A \<^bold>* F g (\<Union>B)"  | 
| 57418 | 423  | 
by (simp add: union_inter_neutral)  | 
| 60758 | 424  | 
with \<open>finite B\<close> \<open>A \<notin> B\<close> show ?case  | 
| 57418 | 425  | 
by (simp add: H)  | 
426  | 
qed  | 
|
427  | 
||
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
428  | 
lemma swap: "F (\<lambda>i. F (g i) B) A = F (\<lambda>j. F (\<lambda>i. g i j) A) B"  | 
| 57418 | 429  | 
unfolding cartesian_product  | 
430  | 
by (rule reindex_bij_witness [where i = "\<lambda>(i, j). (j, i)" and j = "\<lambda>(i, j). (j, i)"]) auto  | 
|
431  | 
||
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
432  | 
lemma swap_restrict:  | 
| 57418 | 433  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow>  | 
434  | 
    F (\<lambda>x. F (g x) {y. y \<in> B \<and> R x y}) A = F (\<lambda>y. F (\<lambda>x. g x y) {x. x \<in> A \<and> R x y}) B"
 | 
|
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
435  | 
by (simp add: inter_filter) (rule swap)  | 
| 57418 | 436  | 
|
437  | 
lemma Plus:  | 
|
438  | 
fixes A :: "'b set" and B :: "'c set"  | 
|
439  | 
assumes fin: "finite A" "finite B"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
440  | 
shows "F g (A <+> B) = F (g \<circ> Inl) A \<^bold>* F (g \<circ> Inr) B"  | 
| 57418 | 441  | 
proof -  | 
442  | 
have "A <+> B = Inl ` A \<union> Inr ` B" by auto  | 
|
| 63654 | 443  | 
moreover from fin have "finite (Inl ` A)" "finite (Inr ` B)" by auto  | 
444  | 
  moreover have "Inl ` A \<inter> Inr ` B = {}" by auto
 | 
|
445  | 
moreover have "inj_on Inl A" "inj_on Inr B" by (auto intro: inj_onI)  | 
|
446  | 
ultimately show ?thesis  | 
|
447  | 
using fin by (simp add: union_disjoint reindex)  | 
|
| 57418 | 448  | 
qed  | 
449  | 
||
| 58195 | 450  | 
lemma same_carrier:  | 
451  | 
assumes "finite C"  | 
|
452  | 
assumes subset: "A \<subseteq> C" "B \<subseteq> C"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
453  | 
assumes trivial: "\<And>a. a \<in> C - A \<Longrightarrow> g a = \<^bold>1" "\<And>b. b \<in> C - B \<Longrightarrow> h b = \<^bold>1"  | 
| 58195 | 454  | 
shows "F g A = F h B \<longleftrightarrow> F g C = F h C"  | 
455  | 
proof -  | 
|
| 63654 | 456  | 
have "finite A" and "finite B" and "finite (C - A)" and "finite (C - B)"  | 
457  | 
using \<open>finite C\<close> subset by (auto elim: finite_subset)  | 
|
| 58195 | 458  | 
from subset have [simp]: "A - (C - A) = A" by auto  | 
459  | 
from subset have [simp]: "B - (C - B) = B" by auto  | 
|
460  | 
from subset have "C = A \<union> (C - A)" by auto  | 
|
461  | 
then have "F g C = F g (A \<union> (C - A))" by simp  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
462  | 
also have "\<dots> = F g (A - (C - A)) \<^bold>* F g (C - A - A) \<^bold>* F g (A \<inter> (C - A))"  | 
| 60758 | 463  | 
using \<open>finite A\<close> \<open>finite (C - A)\<close> by (simp only: union_diff2)  | 
| 63654 | 464  | 
finally have *: "F g C = F g A" using trivial by simp  | 
| 58195 | 465  | 
from subset have "C = B \<union> (C - B)" by auto  | 
466  | 
then have "F h C = F h (B \<union> (C - B))" by simp  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
467  | 
also have "\<dots> = F h (B - (C - B)) \<^bold>* F h (C - B - B) \<^bold>* F h (B \<inter> (C - B))"  | 
| 60758 | 468  | 
using \<open>finite B\<close> \<open>finite (C - B)\<close> by (simp only: union_diff2)  | 
| 63654 | 469  | 
finally have "F h C = F h B"  | 
470  | 
using trivial by simp  | 
|
471  | 
with * show ?thesis by simp  | 
|
| 58195 | 472  | 
qed  | 
473  | 
||
474  | 
lemma same_carrierI:  | 
|
475  | 
assumes "finite C"  | 
|
476  | 
assumes subset: "A \<subseteq> C" "B \<subseteq> C"  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63092 
diff
changeset
 | 
477  | 
assumes trivial: "\<And>a. a \<in> C - A \<Longrightarrow> g a = \<^bold>1" "\<And>b. b \<in> C - B \<Longrightarrow> h b = \<^bold>1"  | 
| 58195 | 478  | 
assumes "F g C = F h C"  | 
479  | 
shows "F g A = F h B"  | 
|
480  | 
using assms same_carrier [of C A B] by simp  | 
|
481  | 
||
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
482  | 
end  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
483  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
484  | 
|
| 60758 | 485  | 
subsection \<open>Generalized summation over a set\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
486  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
487  | 
context comm_monoid_add  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
488  | 
begin  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
489  | 
|
| 64267 | 490  | 
sublocale sum: comm_monoid_set plus 0  | 
491  | 
defines sum = sum.F ..  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
492  | 
|
| 64267 | 493  | 
abbreviation Sum ("\<Sum>_" [1000] 999)
 | 
494  | 
where "\<Sum>A \<equiv> sum (\<lambda>x. x) A"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
495  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
496  | 
end  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
497  | 
|
| 
67268
 
bdf25939a550
new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
 
paulson <lp15@cam.ac.uk> 
parents: 
66936 
diff
changeset
 | 
498  | 
text \<open>Now: lots of fancy syntax. First, @{term "sum (\<lambda>x. e) A"} is written \<open>\<Sum>x\<in>A. e\<close>.\<close>
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
499  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
500  | 
syntax (ASCII)  | 
| 
67268
 
bdf25939a550
new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
 
paulson <lp15@cam.ac.uk> 
parents: 
66936 
diff
changeset
 | 
501  | 
  "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add"  ("(3SUM (_/:_)./ _)" [0, 51, 10] 10)
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
502  | 
syntax  | 
| 
67268
 
bdf25939a550
new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
 
paulson <lp15@cam.ac.uk> 
parents: 
66936 
diff
changeset
 | 
503  | 
  "_sum" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b::comm_monoid_add"  ("(2\<Sum>(_/\<in>_)./ _)" [0, 51, 10] 10)
 | 
| 61799 | 504  | 
translations \<comment> \<open>Beware of argument permutation!\<close>  | 
| 64267 | 505  | 
"\<Sum>i\<in>A. b" \<rightleftharpoons> "CONST sum (\<lambda>i. b) A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
506  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
507  | 
text \<open>Instead of @{term"\<Sum>x\<in>{x. P}. e"} we introduce the shorter \<open>\<Sum>x|P. e\<close>.\<close>
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
508  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
509  | 
syntax (ASCII)  | 
| 64267 | 510  | 
  "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(3SUM _ |/ _./ _)" [0, 0, 10] 10)
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
511  | 
syntax  | 
| 64267 | 512  | 
  "_qsum" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Sum>_ | (_)./ _)" [0, 0, 10] 10)
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
513  | 
translations  | 
| 64267 | 514  | 
  "\<Sum>x|P. t" => "CONST sum (\<lambda>x. t) {x. P}"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
515  | 
|
| 60758 | 516  | 
print_translation \<open>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
517  | 
let  | 
| 64267 | 518  | 
  fun sum_tr' [Abs (x, Tx, t), Const (@{const_syntax Collect}, _) $ Abs (y, Ty, P)] =
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
519  | 
if x <> y then raise Match  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
520  | 
else  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
521  | 
let  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
522  | 
val x' = Syntax_Trans.mark_bound_body (x, Tx);  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
523  | 
val t' = subst_bound (x', t);  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
524  | 
val P' = subst_bound (x', P);  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
525  | 
in  | 
| 64267 | 526  | 
            Syntax.const @{syntax_const "_qsum"} $ Syntax_Trans.mark_bound_abs (x, Tx) $ P' $ t'
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
527  | 
end  | 
| 64267 | 528  | 
| sum_tr' _ = raise Match;  | 
529  | 
in [(@{const_syntax sum}, K sum_tr')] end
 | 
|
| 60758 | 530  | 
\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
531  | 
|
| 63654 | 532  | 
(* TODO generalization candidates *)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
533  | 
|
| 64267 | 534  | 
lemma (in comm_monoid_add) sum_image_gen:  | 
| 63654 | 535  | 
assumes fin: "finite S"  | 
| 64267 | 536  | 
  shows "sum g S = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
 | 
| 63654 | 537  | 
proof -  | 
538  | 
  have "{y. y\<in> f`S \<and> f x = y} = {f x}" if "x \<in> S" for x
 | 
|
539  | 
using that by auto  | 
|
| 64267 | 540  | 
  then have "sum g S = sum (\<lambda>x. sum (\<lambda>y. g x) {y. y\<in> f`S \<and> f x = y}) S"
 | 
| 57418 | 541  | 
by simp  | 
| 64267 | 542  | 
  also have "\<dots> = sum (\<lambda>y. sum g {x. x \<in> S \<and> f x = y}) (f ` S)"
 | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
543  | 
by (rule sum.swap_restrict [OF fin finite_imageI [OF fin]])  | 
| 57418 | 544  | 
finally show ?thesis .  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
545  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
546  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
547  | 
|
| 60758 | 548  | 
subsubsection \<open>Properties in more restricted classes of structures\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
549  | 
|
| 64267 | 550  | 
lemma sum_Un:  | 
551  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow> sum f (A \<union> B) = sum f A + sum f B - sum f (A \<inter> B)"  | 
|
| 63654 | 552  | 
for f :: "'b \<Rightarrow> 'a::ab_group_add"  | 
| 64267 | 553  | 
by (subst sum.union_inter [symmetric]) (auto simp add: algebra_simps)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
554  | 
|
| 64267 | 555  | 
lemma sum_Un2:  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
556  | 
assumes "finite (A \<union> B)"  | 
| 64267 | 557  | 
shows "sum f (A \<union> B) = sum f (A - B) + sum f (B - A) + sum f (A \<inter> B)"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
558  | 
proof -  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
559  | 
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
 | 
560  | 
by auto  | 
| 63654 | 561  | 
with assms show ?thesis  | 
| 64267 | 562  | 
by simp (subst sum.union_disjoint, auto)+  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
563  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
564  | 
|
| 64267 | 565  | 
lemma sum_diff1:  | 
| 63654 | 566  | 
fixes f :: "'b \<Rightarrow> 'a::ab_group_add"  | 
567  | 
assumes "finite A"  | 
|
| 64267 | 568  | 
  shows "sum f (A - {a}) = (if a \<in> A then sum f A - f a else sum f A)"
 | 
| 63654 | 569  | 
using assms by induct (auto simp: insert_Diff_if)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
570  | 
|
| 64267 | 571  | 
lemma sum_diff:  | 
| 63654 | 572  | 
fixes f :: "'b \<Rightarrow> 'a::ab_group_add"  | 
573  | 
assumes "finite A" "B \<subseteq> A"  | 
|
| 64267 | 574  | 
shows "sum f (A - B) = sum f A - sum f B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
575  | 
proof -  | 
| 63654 | 576  | 
from assms(2,1) have "finite B" by (rule finite_subset)  | 
577  | 
from this \<open>B \<subseteq> A\<close>  | 
|
578  | 
show ?thesis  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
579  | 
proof induct  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
580  | 
case empty  | 
| 63654 | 581  | 
thus ?case by simp  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
582  | 
next  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
583  | 
case (insert x F)  | 
| 63654 | 584  | 
with \<open>finite A\<close> \<open>finite B\<close> show ?case  | 
| 64267 | 585  | 
by (simp add: Diff_insert[where a=x and B=F] sum_diff1 insert_absorb)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
586  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
587  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
588  | 
|
| 64267 | 589  | 
lemma (in ordered_comm_monoid_add) sum_mono:  | 
| 63915 | 590  | 
"(\<And>i. i\<in>K \<Longrightarrow> f i \<le> g i) \<Longrightarrow> (\<Sum>i\<in>K. f i) \<le> (\<Sum>i\<in>K. g i)"  | 
591  | 
by (induct K rule: infinite_finite_induct) (use add_mono in auto)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
592  | 
|
| 64267 | 593  | 
lemma (in strict_ordered_comm_monoid_add) sum_strict_mono:  | 
| 63654 | 594  | 
  assumes "finite A" "A \<noteq> {}"
 | 
595  | 
and "\<And>x. x \<in> A \<Longrightarrow> f x < g x"  | 
|
| 64267 | 596  | 
shows "sum f A < sum g A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
597  | 
using assms  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
598  | 
proof (induct rule: finite_ne_induct)  | 
| 63654 | 599  | 
case singleton  | 
600  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
601  | 
next  | 
| 63654 | 602  | 
case insert  | 
603  | 
then show ?case by (auto simp: add_strict_mono)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
604  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
605  | 
|
| 64267 | 606  | 
lemma sum_strict_mono_ex1:  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
607  | 
fixes f g :: "'i \<Rightarrow> 'a::ordered_cancel_comm_monoid_add"  | 
| 63654 | 608  | 
assumes "finite A"  | 
609  | 
and "\<forall>x\<in>A. f x \<le> g x"  | 
|
610  | 
and "\<exists>a\<in>A. f a < g a"  | 
|
| 64267 | 611  | 
shows "sum f A < sum g A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
612  | 
proof-  | 
| 63654 | 613  | 
from assms(3) obtain a where a: "a \<in> A" "f a < g a" by blast  | 
| 64267 | 614  | 
  have "sum f A = sum f ((A - {a}) \<union> {a})"
 | 
| 63654 | 615  | 
by(simp add: insert_absorb[OF \<open>a \<in> A\<close>])  | 
| 64267 | 616  | 
  also have "\<dots> = sum f (A - {a}) + sum f {a}"
 | 
617  | 
using \<open>finite A\<close> by(subst sum.union_disjoint) auto  | 
|
618  | 
  also have "sum f (A - {a}) \<le> sum g (A - {a})"
 | 
|
619  | 
by (rule sum_mono) (simp add: assms(2))  | 
|
620  | 
  also from a have "sum f {a} < sum g {a}" by simp
 | 
|
621  | 
  also have "sum g (A - {a}) + sum g {a} = sum g((A - {a}) \<union> {a})"
 | 
|
622  | 
using \<open>finite A\<close> by (subst sum.union_disjoint[symmetric]) auto  | 
|
623  | 
also have "\<dots> = sum g A" by (simp add: insert_absorb[OF \<open>a \<in> A\<close>])  | 
|
| 63654 | 624  | 
finally show ?thesis  | 
625  | 
by (auto simp add: add_right_mono add_strict_left_mono)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
626  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
627  | 
|
| 64267 | 628  | 
lemma sum_mono_inv:  | 
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
629  | 
fixes f g :: "'i \<Rightarrow> 'a :: ordered_cancel_comm_monoid_add"  | 
| 64267 | 630  | 
assumes eq: "sum f I = sum g I"  | 
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
631  | 
assumes le: "\<And>i. i \<in> I \<Longrightarrow> f i \<le> g i"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
632  | 
assumes i: "i \<in> I"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
633  | 
assumes I: "finite I"  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
634  | 
shows "f i = g i"  | 
| 63654 | 635  | 
proof (rule ccontr)  | 
636  | 
assume "\<not> ?thesis"  | 
|
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
637  | 
with le[OF i] have "f i < g i" by simp  | 
| 63654 | 638  | 
with i have "\<exists>i\<in>I. f i < g i" ..  | 
| 64267 | 639  | 
from sum_strict_mono_ex1[OF I _ this] le have "sum f I < sum g I"  | 
| 63654 | 640  | 
by blast  | 
| 
63561
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
641  | 
with eq show False by simp  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
642  | 
qed  | 
| 
 
fba08009ff3e
add lemmas contributed by Peter Gammie
 
Andreas Lochbihler 
parents: 
63357 
diff
changeset
 | 
643  | 
|
| 64267 | 644  | 
lemma member_le_sum:  | 
| 63938 | 645  | 
  fixes f :: "_ \<Rightarrow> 'b::{semiring_1, ordered_comm_monoid_add}"
 | 
| 
66112
 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
66089 
diff
changeset
 | 
646  | 
assumes "i \<in> A"  | 
| 
 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 
paulson <lp15@cam.ac.uk> 
parents: 
66089 
diff
changeset
 | 
647  | 
    and le: "\<And>x. x \<in> A - {i} \<Longrightarrow> 0 \<le> f x"
 | 
| 63938 | 648  | 
and "finite A"  | 
| 64267 | 649  | 
shows "f i \<le> sum f A"  | 
| 63938 | 650  | 
proof -  | 
| 64267 | 651  | 
  have "f i \<le> sum f (A \<inter> {i})"
 | 
| 63938 | 652  | 
by (simp add: assms)  | 
653  | 
  also have "... = (\<Sum>x\<in>A. if x \<in> {i} then f x else 0)"
 | 
|
| 64267 | 654  | 
using assms sum.inter_restrict by blast  | 
655  | 
also have "... \<le> sum f A"  | 
|
656  | 
apply (rule sum_mono)  | 
|
| 63938 | 657  | 
apply (auto simp: le)  | 
658  | 
done  | 
|
659  | 
finally show ?thesis .  | 
|
660  | 
qed  | 
|
661  | 
||
| 64267 | 662  | 
lemma sum_negf: "(\<Sum>x\<in>A. - f x) = - (\<Sum>x\<in>A. f x)"  | 
| 63654 | 663  | 
for f :: "'b \<Rightarrow> 'a::ab_group_add"  | 
| 63915 | 664  | 
by (induct A rule: infinite_finite_induct) auto  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
665  | 
|
| 64267 | 666  | 
lemma sum_subtractf: "(\<Sum>x\<in>A. f x - g x) = (\<Sum>x\<in>A. f x) - (\<Sum>x\<in>A. g x)"  | 
| 63654 | 667  | 
for f g :: "'b \<Rightarrow>'a::ab_group_add"  | 
| 64267 | 668  | 
using sum.distrib [of f "- g" A] by (simp add: sum_negf)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
669  | 
|
| 64267 | 670  | 
lemma sum_subtractf_nat:  | 
| 63654 | 671  | 
"(\<And>x. x \<in> A \<Longrightarrow> g x \<le> f x) \<Longrightarrow> (\<Sum>x\<in>A. f x - g x) = (\<Sum>x\<in>A. f x) - (\<Sum>x\<in>A. g x)"  | 
672  | 
for f g :: "'a \<Rightarrow> nat"  | 
|
| 64267 | 673  | 
by (induct A rule: infinite_finite_induct) (auto simp: sum_mono)  | 
| 
59416
 
fde2659085e1
generalized sum_diff_distrib to setsum_subtractf_nat
 
hoelzl 
parents: 
59010 
diff
changeset
 | 
674  | 
|
| 63654 | 675  | 
context ordered_comm_monoid_add  | 
676  | 
begin  | 
|
677  | 
||
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
678  | 
lemma sum_nonneg: "(\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> sum f A"  | 
| 63915 | 679  | 
proof (induct A rule: infinite_finite_induct)  | 
680  | 
case infinite  | 
|
681  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
682  | 
next  | 
| 63915 | 683  | 
case empty  | 
684  | 
then show ?case by simp  | 
|
685  | 
next  | 
|
686  | 
case (insert x F)  | 
|
| 64267 | 687  | 
then have "0 + 0 \<le> f x + sum f F" by (blast intro: add_mono)  | 
| 63915 | 688  | 
with insert show ?case by simp  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
689  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
690  | 
|
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
691  | 
lemma sum_nonpos: "(\<And>x. x \<in> A \<Longrightarrow> f x \<le> 0) \<Longrightarrow> sum f A \<le> 0"  | 
| 63915 | 692  | 
proof (induct A rule: infinite_finite_induct)  | 
693  | 
case infinite  | 
|
694  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
695  | 
next  | 
| 63915 | 696  | 
case empty  | 
697  | 
then show ?case by simp  | 
|
698  | 
next  | 
|
699  | 
case (insert x F)  | 
|
| 64267 | 700  | 
then have "f x + sum f F \<le> 0 + 0" by (blast intro: add_mono)  | 
| 63915 | 701  | 
with insert show ?case by simp  | 
| 
54744
 
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  | 
|
| 64267 | 704  | 
lemma sum_nonneg_eq_0_iff:  | 
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
705  | 
"finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> sum f A = 0 \<longleftrightarrow> (\<forall>x\<in>A. f x = 0)"  | 
| 64267 | 706  | 
by (induct set: finite) (simp_all add: add_nonneg_eq_0_iff sum_nonneg)  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
707  | 
|
| 64267 | 708  | 
lemma sum_nonneg_0:  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
709  | 
"finite s \<Longrightarrow> (\<And>i. i \<in> s \<Longrightarrow> f i \<ge> 0) \<Longrightarrow> (\<Sum> i \<in> s. f i) = 0 \<Longrightarrow> i \<in> s \<Longrightarrow> f i = 0"  | 
| 64267 | 710  | 
by (simp add: sum_nonneg_eq_0_iff)  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
711  | 
|
| 64267 | 712  | 
lemma sum_nonneg_leq_bound:  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
713  | 
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
 | 
714  | 
shows "f i \<le> B"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
715  | 
proof -  | 
| 63654 | 716  | 
  from assms have "f i \<le> f i + (\<Sum>i \<in> s - {i}. f i)"
 | 
| 64267 | 717  | 
by (intro add_increasing2 sum_nonneg) auto  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
718  | 
also have "\<dots> = B"  | 
| 64267 | 719  | 
using sum.remove[of s i f] assms by simp  | 
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
720  | 
finally show ?thesis by auto  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
721  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
722  | 
|
| 64267 | 723  | 
lemma sum_mono2:  | 
| 63654 | 724  | 
assumes fin: "finite B"  | 
725  | 
and sub: "A \<subseteq> B"  | 
|
726  | 
and nn: "\<And>b. b \<in> B-A \<Longrightarrow> 0 \<le> f b"  | 
|
| 64267 | 727  | 
shows "sum f A \<le> sum f B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
728  | 
proof -  | 
| 64267 | 729  | 
have "sum f A \<le> sum f A + sum f (B-A)"  | 
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
730  | 
by (auto intro: add_increasing2 [OF sum_nonneg] nn)  | 
| 64267 | 731  | 
also from fin finite_subset[OF sub fin] have "\<dots> = sum f (A \<union> (B-A))"  | 
732  | 
by (simp add: sum.union_disjoint del: Un_Diff_cancel)  | 
|
| 63654 | 733  | 
also from sub have "A \<union> (B-A) = B" by blast  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
734  | 
finally show ?thesis .  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
735  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
736  | 
|
| 64267 | 737  | 
lemma sum_le_included:  | 
| 57418 | 738  | 
assumes "finite s" "finite t"  | 
739  | 
and "\<forall>y\<in>t. 0 \<le> g y" "(\<forall>x\<in>s. \<exists>y\<in>t. i y = x \<and> f x \<le> g y)"  | 
|
| 64267 | 740  | 
shows "sum f s \<le> sum g t"  | 
| 57418 | 741  | 
proof -  | 
| 64267 | 742  | 
  have "sum f s \<le> sum (\<lambda>y. sum g {x. x\<in>t \<and> i x = y}) s"
 | 
743  | 
proof (rule sum_mono)  | 
|
| 63654 | 744  | 
fix y  | 
745  | 
assume "y \<in> s"  | 
|
| 57418 | 746  | 
with assms obtain z where z: "z \<in> t" "y = i z" "f y \<le> g z" by auto  | 
| 64267 | 747  | 
    with assms show "f y \<le> sum g {x \<in> t. i x = y}" (is "?A y \<le> ?B y")
 | 
748  | 
      using order_trans[of "?A (i z)" "sum g {z}" "?B (i z)", intro]
 | 
|
749  | 
by (auto intro!: sum_mono2)  | 
|
| 57418 | 750  | 
qed  | 
| 64267 | 751  | 
  also have "\<dots> \<le> sum (\<lambda>y. sum g {x. x\<in>t \<and> i x = y}) (i ` t)"
 | 
752  | 
using assms(2-4) by (auto intro!: sum_mono2 sum_nonneg)  | 
|
753  | 
also have "\<dots> \<le> sum g t"  | 
|
754  | 
using assms by (auto simp: sum_image_gen[symmetric])  | 
|
| 57418 | 755  | 
finally show ?thesis .  | 
756  | 
qed  | 
|
757  | 
||
| 63654 | 758  | 
end  | 
759  | 
||
| 64267 | 760  | 
lemma (in canonically_ordered_monoid_add) sum_eq_0_iff [simp]:  | 
761  | 
"finite F \<Longrightarrow> (sum f F = 0) = (\<forall>a\<in>F. f a = 0)"  | 
|
762  | 
by (intro ballI sum_nonneg_eq_0_iff zero_le)  | 
|
| 
62376
 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 
hoelzl 
parents: 
61955 
diff
changeset
 | 
763  | 
|
| 66936 | 764  | 
context semiring_0  | 
765  | 
begin  | 
|
766  | 
||
767  | 
lemma sum_distrib_left: "r * sum f A = (\<Sum>n\<in>A. r * f n)"  | 
|
768  | 
by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
769  | 
|
| 64267 | 770  | 
lemma sum_distrib_right: "sum f A * r = (\<Sum>n\<in>A. f n * r)"  | 
| 66936 | 771  | 
by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps)  | 
772  | 
||
773  | 
end  | 
|
| 63654 | 774  | 
|
| 64267 | 775  | 
lemma sum_divide_distrib: "sum f A / r = (\<Sum>n\<in>A. f n / r)"  | 
| 63654 | 776  | 
for r :: "'a::field"  | 
| 63915 | 777  | 
proof (induct A rule: infinite_finite_induct)  | 
778  | 
case infinite  | 
|
779  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
780  | 
next  | 
| 63915 | 781  | 
case empty  | 
782  | 
then show ?case by simp  | 
|
783  | 
next  | 
|
784  | 
case insert  | 
|
785  | 
then show ?case by (simp add: add_divide_distrib)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
786  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
787  | 
|
| 64267 | 788  | 
lemma sum_abs[iff]: "\<bar>sum f A\<bar> \<le> sum (\<lambda>i. \<bar>f i\<bar>) A"  | 
| 63654 | 789  | 
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"  | 
| 63915 | 790  | 
proof (induct A rule: infinite_finite_induct)  | 
791  | 
case infinite  | 
|
792  | 
then show ?case by simp  | 
|
| 63654 | 793  | 
next  | 
| 63915 | 794  | 
case empty  | 
795  | 
then show ?case by simp  | 
|
796  | 
next  | 
|
797  | 
case insert  | 
|
798  | 
then show ?case by (auto intro: abs_triangle_ineq order_trans)  | 
|
| 63654 | 799  | 
qed  | 
800  | 
||
| 64267 | 801  | 
lemma sum_abs_ge_zero[iff]: "0 \<le> sum (\<lambda>i. \<bar>f i\<bar>) A"  | 
| 63654 | 802  | 
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"  | 
| 64267 | 803  | 
by (simp add: sum_nonneg)  | 
| 63654 | 804  | 
|
| 64267 | 805  | 
lemma abs_sum_abs[simp]: "\<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar> = (\<Sum>a\<in>A. \<bar>f a\<bar>)"  | 
| 63654 | 806  | 
for f :: "'a \<Rightarrow> 'b::ordered_ab_group_add_abs"  | 
| 63915 | 807  | 
proof (induct A rule: infinite_finite_induct)  | 
808  | 
case infinite  | 
|
809  | 
then show ?case by simp  | 
|
810  | 
next  | 
|
811  | 
case empty  | 
|
812  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
813  | 
next  | 
| 63915 | 814  | 
case (insert a A)  | 
815  | 
then have "\<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  | 
|
816  | 
also from insert have "\<dots> = \<bar>\<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>\<bar>" by simp  | 
|
817  | 
also have "\<dots> = \<bar>f a\<bar> + \<bar>\<Sum>a\<in>A. \<bar>f a\<bar>\<bar>" by (simp del: abs_of_nonneg)  | 
|
818  | 
also from insert have "\<dots> = (\<Sum>a\<in>insert a A. \<bar>f a\<bar>)" by simp  | 
|
819  | 
finally show ?case .  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
820  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
821  | 
|
| 64267 | 822  | 
lemma sum_diff1_ring:  | 
| 63654 | 823  | 
fixes f :: "'b \<Rightarrow> 'a::ring"  | 
824  | 
assumes "finite A" "a \<in> A"  | 
|
| 64267 | 825  | 
  shows "sum f (A - {a}) = sum f A - (f a)"
 | 
826  | 
unfolding sum.remove [OF assms] by auto  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
827  | 
|
| 64267 | 828  | 
lemma sum_product:  | 
| 63654 | 829  | 
fixes f :: "'a \<Rightarrow> 'b::semiring_0"  | 
| 64267 | 830  | 
shows "sum f A * sum g B = (\<Sum>i\<in>A. \<Sum>j\<in>B. f i * g j)"  | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
831  | 
by (simp add: sum_distrib_left sum_distrib_right) (rule sum.swap)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
832  | 
|
| 64267 | 833  | 
lemma sum_mult_sum_if_inj:  | 
| 63654 | 834  | 
fixes f :: "'a \<Rightarrow> 'b::semiring_0"  | 
835  | 
shows "inj_on (\<lambda>(a, b). f a * g b) (A \<times> B) \<Longrightarrow>  | 
|
| 64267 | 836  | 
    sum f A * sum g B = sum id {f a * g b |a b. a \<in> A \<and> b \<in> B}"
 | 
837  | 
by(auto simp: sum_product sum.cartesian_product intro!: sum.reindex_cong[symmetric])  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
838  | 
|
| 64267 | 839  | 
lemma sum_SucD: "sum f A = Suc n \<Longrightarrow> \<exists>a\<in>A. 0 < f a"  | 
| 63915 | 840  | 
by (induct A rule: infinite_finite_induct) auto  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
841  | 
|
| 64267 | 842  | 
lemma sum_eq_Suc0_iff:  | 
843  | 
"finite A \<Longrightarrow> sum f A = Suc 0 \<longleftrightarrow> (\<exists>a\<in>A. f a = Suc 0 \<and> (\<forall>b\<in>A. a \<noteq> b \<longrightarrow> f b = 0))"  | 
|
| 63915 | 844  | 
by (induct A rule: finite_induct) (auto simp add: add_is_1)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
845  | 
|
| 64267 | 846  | 
lemmas sum_eq_1_iff = sum_eq_Suc0_iff[simplified One_nat_def[symmetric]]  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
847  | 
|
| 64267 | 848  | 
lemma sum_Un_nat:  | 
849  | 
"finite A \<Longrightarrow> finite B \<Longrightarrow> sum f (A \<union> B) = sum f A + sum f B - sum f (A \<inter> B)"  | 
|
| 63654 | 850  | 
for f :: "'a \<Rightarrow> nat"  | 
| 61799 | 851  | 
\<comment> \<open>For the natural numbers, we have subtraction.\<close>  | 
| 64267 | 852  | 
by (subst sum.union_inter [symmetric]) (auto simp: algebra_simps)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
853  | 
|
| 64267 | 854  | 
lemma sum_diff1_nat: "sum f (A - {a}) = (if a \<in> A then sum f A - f a else sum f A)"
 | 
| 63654 | 855  | 
for f :: "'a \<Rightarrow> nat"  | 
| 63915 | 856  | 
proof (induct A rule: infinite_finite_induct)  | 
857  | 
case infinite  | 
|
858  | 
then show ?case by simp  | 
|
859  | 
next  | 
|
860  | 
case empty  | 
|
861  | 
then show ?case by simp  | 
|
862  | 
next  | 
|
863  | 
case insert  | 
|
864  | 
then show ?case  | 
|
865  | 
apply (auto simp: insert_Diff_if)  | 
|
| 63654 | 866  | 
apply (drule mk_disjoint_insert)  | 
867  | 
apply auto  | 
|
868  | 
done  | 
|
869  | 
qed  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
870  | 
|
| 64267 | 871  | 
lemma sum_diff_nat:  | 
| 63654 | 872  | 
fixes f :: "'a \<Rightarrow> nat"  | 
873  | 
assumes "finite B" and "B \<subseteq> A"  | 
|
| 64267 | 874  | 
shows "sum f (A - B) = sum f A - sum f B"  | 
| 63654 | 875  | 
using assms  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
876  | 
proof induct  | 
| 63654 | 877  | 
case empty  | 
878  | 
then show ?case by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
879  | 
next  | 
| 63654 | 880  | 
case (insert x F)  | 
| 64267 | 881  | 
note IH = \<open>F \<subseteq> A \<Longrightarrow> sum f (A - F) = sum f A - sum f F\<close>  | 
| 63654 | 882  | 
from \<open>x \<notin> F\<close> \<open>insert x F \<subseteq> A\<close> have "x \<in> A - F" by simp  | 
| 64267 | 883  | 
  then have A: "sum f ((A - F) - {x}) = sum f (A - F) - f x"
 | 
884  | 
by (simp add: sum_diff1_nat)  | 
|
| 63654 | 885  | 
from \<open>insert x F \<subseteq> A\<close> have "F \<subseteq> A" by simp  | 
| 64267 | 886  | 
with IH have "sum f (A - F) = sum f A - sum f F" by simp  | 
887  | 
  with A have B: "sum f ((A - F) - {x}) = sum f A - sum f F - f x"
 | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
888  | 
by simp  | 
| 63654 | 889  | 
  from \<open>x \<notin> F\<close> have "A - insert x F = (A - F) - {x}" by auto
 | 
| 64267 | 890  | 
with B have C: "sum f (A - insert x F) = sum f A - sum f F - f x"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
891  | 
by simp  | 
| 64267 | 892  | 
from \<open>finite F\<close> \<open>x \<notin> F\<close> have "sum f (insert x F) = sum f F + f x"  | 
| 63654 | 893  | 
by simp  | 
| 64267 | 894  | 
with C have "sum f (A - insert x F) = sum f A - sum f (insert x F)"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
895  | 
by simp  | 
| 63654 | 896  | 
then show ?case by simp  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
897  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
898  | 
|
| 64267 | 899  | 
lemma sum_comp_morphism:  | 
900  | 
"h 0 = 0 \<Longrightarrow> (\<And>x y. h (x + y) = h x + h y) \<Longrightarrow> sum (h \<circ> g) A = h (sum g A)"  | 
|
| 63915 | 901  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
902  | 
|
| 64267 | 903  | 
lemma (in comm_semiring_1) dvd_sum: "(\<And>a. a \<in> A \<Longrightarrow> d dvd f a) \<Longrightarrow> d dvd sum f A"  | 
| 59010 | 904  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
905  | 
||
| 64267 | 906  | 
lemma (in ordered_comm_monoid_add) sum_pos:  | 
907  | 
  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> 0 < f i) \<Longrightarrow> 0 < sum f I"
 | 
|
| 
62377
 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 
hoelzl 
parents: 
62376 
diff
changeset
 | 
908  | 
by (induct I rule: finite_ne_induct) (auto intro: add_pos_pos)  | 
| 
 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 
hoelzl 
parents: 
62376 
diff
changeset
 | 
909  | 
|
| 64267 | 910  | 
lemma (in ordered_comm_monoid_add) sum_pos2:  | 
| 
62377
 
ace69956d018
moved more proofs to ordered_comm_monoid_add; introduced strict_ordered_ab_semigroup/comm_monoid_add
 
hoelzl 
parents: 
62376 
diff
changeset
 | 
911  | 
assumes I: "finite I" "i \<in> I" "0 < f i" "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> f i"  | 
| 64267 | 912  | 
shows "0 < sum f I"  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
913  | 
proof -  | 
| 64267 | 914  | 
  have "0 < f i + sum f (I - {i})"
 | 
915  | 
using assms by (intro add_pos_nonneg sum_nonneg) auto  | 
|
916  | 
also have "\<dots> = sum f I"  | 
|
917  | 
using assms by (simp add: sum.remove)  | 
|
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
918  | 
finally show ?thesis .  | 
| 
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
919  | 
qed  | 
| 
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
920  | 
|
| 64267 | 921  | 
lemma sum_cong_Suc:  | 
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61378 
diff
changeset
 | 
922  | 
assumes "0 \<notin> A" "\<And>x. Suc x \<in> A \<Longrightarrow> f (Suc x) = g (Suc x)"  | 
| 64267 | 923  | 
shows "sum f A = sum g A"  | 
924  | 
proof (rule sum.cong)  | 
|
| 63654 | 925  | 
fix x  | 
926  | 
assume "x \<in> A"  | 
|
927  | 
with assms(1) show "f x = g x"  | 
|
928  | 
by (cases x) (auto intro!: assms(2))  | 
|
| 
61524
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61378 
diff
changeset
 | 
929  | 
qed simp_all  | 
| 
 
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
 
eberlm 
parents: 
61378 
diff
changeset
 | 
930  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
931  | 
|
| 64267 | 932  | 
subsubsection \<open>Cardinality as special case of @{const sum}\<close>
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
933  | 
|
| 64267 | 934  | 
lemma card_eq_sum: "card A = sum (\<lambda>x. 1) A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
935  | 
proof -  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
936  | 
have "plus \<circ> (\<lambda>_. Suc 0) = (\<lambda>_. Suc)"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
937  | 
by (simp add: fun_eq_iff)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
938  | 
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
 | 
939  | 
by (rule arg_cong)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
940  | 
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
 | 
941  | 
by (blast intro: fun_cong)  | 
| 63654 | 942  | 
then show ?thesis  | 
| 64267 | 943  | 
by (simp add: card.eq_fold sum.eq_fold)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
944  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
945  | 
|
| 66936 | 946  | 
context semiring_1  | 
947  | 
begin  | 
|
948  | 
||
949  | 
lemma sum_constant [simp]:  | 
|
950  | 
"(\<Sum>x \<in> A. y) = of_nat (card A) * y"  | 
|
951  | 
by (induct A rule: infinite_finite_induct) (simp_all add: algebra_simps)  | 
|
952  | 
||
953  | 
end  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
954  | 
|
| 64267 | 955  | 
lemma sum_Suc: "sum (\<lambda>x. Suc(f x)) A = sum f A + card A"  | 
956  | 
using sum.distrib[of f "\<lambda>_. 1" A] by simp  | 
|
| 58349 | 957  | 
|
| 64267 | 958  | 
lemma sum_bounded_above:  | 
| 63654 | 959  | 
  fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}"
 | 
960  | 
assumes le: "\<And>i. i\<in>A \<Longrightarrow> f i \<le> K"  | 
|
| 64267 | 961  | 
shows "sum f A \<le> of_nat (card A) * K"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
962  | 
proof (cases "finite A")  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
963  | 
case True  | 
| 63654 | 964  | 
then show ?thesis  | 
| 64267 | 965  | 
using le sum_mono[where K=A and g = "\<lambda>x. K"] by simp  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
966  | 
next  | 
| 63654 | 967  | 
case False  | 
968  | 
then show ?thesis by simp  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
969  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
970  | 
|
| 64267 | 971  | 
lemma sum_bounded_above_strict:  | 
| 63654 | 972  | 
  fixes K :: "'a::{ordered_cancel_comm_monoid_add,semiring_1}"
 | 
973  | 
assumes "\<And>i. i\<in>A \<Longrightarrow> f i < K" "card A > 0"  | 
|
| 64267 | 974  | 
shows "sum f A < of_nat (card A) * K"  | 
975  | 
using assms sum_strict_mono[where A=A and g = "\<lambda>x. K"]  | 
|
| 63654 | 976  | 
by (simp add: card_gt_0_iff)  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
977  | 
|
| 64267 | 978  | 
lemma sum_bounded_below:  | 
| 63654 | 979  | 
  fixes K :: "'a::{semiring_1,ordered_comm_monoid_add}"
 | 
980  | 
assumes le: "\<And>i. i\<in>A \<Longrightarrow> K \<le> f i"  | 
|
| 64267 | 981  | 
shows "of_nat (card A) * K \<le> sum f A"  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
982  | 
proof (cases "finite A")  | 
| 
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
983  | 
case True  | 
| 63654 | 984  | 
then show ?thesis  | 
| 64267 | 985  | 
using le sum_mono[where K=A and f = "\<lambda>x. K"] by simp  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
986  | 
next  | 
| 63654 | 987  | 
case False  | 
988  | 
then show ?thesis by simp  | 
|
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
989  | 
qed  | 
| 
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
990  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
991  | 
lemma card_UN_disjoint:  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
992  | 
assumes "finite I" and "\<forall>i\<in>I. finite (A i)"  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
993  | 
    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
 | 
994  | 
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
 | 
995  | 
proof -  | 
| 63654 | 996  | 
have "(\<Sum>i\<in>I. card (A i)) = (\<Sum>i\<in>I. \<Sum>x\<in>A i. 1)"  | 
997  | 
by simp  | 
|
998  | 
with assms show ?thesis  | 
|
| 64267 | 999  | 
by (simp add: card_eq_sum sum.UNION_disjoint del: sum_constant)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1000  | 
qed  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1001  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1002  | 
lemma card_Union_disjoint:  | 
| 63654 | 1003  | 
  "finite C \<Longrightarrow> \<forall>A\<in>C. finite A \<Longrightarrow> \<forall>A\<in>C. \<forall>B\<in>C. A \<noteq> B \<longrightarrow> A \<inter> B = {} \<Longrightarrow>
 | 
| 64267 | 1004  | 
card (\<Union>C) = sum card C"  | 
| 63654 | 1005  | 
by (frule card_UN_disjoint [of C id]) simp_all  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1006  | 
|
| 64267 | 1007  | 
lemma sum_multicount_gen:  | 
| 57418 | 1008  | 
  assumes "finite s" "finite t" "\<forall>j\<in>t. (card {i\<in>s. R i j} = k j)"
 | 
| 64267 | 1009  | 
  shows "sum (\<lambda>i. (card {j\<in>t. R i j})) s = sum k t"
 | 
| 63654 | 1010  | 
(is "?l = ?r")  | 
| 57418 | 1011  | 
proof-  | 
| 64267 | 1012  | 
  have "?l = sum (\<lambda>i. sum (\<lambda>x.1) {j\<in>t. R i j}) s"
 | 
| 63654 | 1013  | 
by auto  | 
1014  | 
also have "\<dots> = ?r"  | 
|
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
66364 
diff
changeset
 | 
1015  | 
unfolding sum.swap_restrict [OF assms(1-2)]  | 
| 57418 | 1016  | 
using assms(3) by auto  | 
1017  | 
finally show ?thesis .  | 
|
1018  | 
qed  | 
|
1019  | 
||
| 64267 | 1020  | 
lemma sum_multicount:  | 
| 57418 | 1021  | 
  assumes "finite S" "finite T" "\<forall>j\<in>T. (card {i\<in>S. R i j} = k)"
 | 
| 64267 | 1022  | 
  shows "sum (\<lambda>i. card {j\<in>T. R i j}) S = k * card T" (is "?l = ?r")
 | 
| 57418 | 1023  | 
proof-  | 
| 64267 | 1024  | 
have "?l = sum (\<lambda>i. k) T"  | 
1025  | 
by (rule sum_multicount_gen) (auto simp: assms)  | 
|
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57418 
diff
changeset
 | 
1026  | 
also have "\<dots> = ?r" by (simp add: mult.commute)  | 
| 57418 | 1027  | 
finally show ?thesis by auto  | 
1028  | 
qed  | 
|
1029  | 
||
| 63654 | 1030  | 
|
| 60758 | 1031  | 
subsubsection \<open>Cardinality of products\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1032  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1033  | 
lemma card_SigmaI [simp]:  | 
| 63654 | 1034  | 
"finite A \<Longrightarrow> \<forall>a\<in>A. finite (B a) \<Longrightarrow> card (SIGMA x: A. B x) = (\<Sum>a\<in>A. card (B a))"  | 
| 64267 | 1035  | 
by (simp add: card_eq_sum sum.Sigma del: sum_constant)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1036  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1037  | 
(*  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1038  | 
lemma SigmaI_insert: "y \<notin> A ==>  | 
| 61943 | 1039  | 
  (SIGMA x:(insert y A). B x) = (({y} \<times> (B y)) \<union> (SIGMA x: A. B x))"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1040  | 
by auto  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1041  | 
*)  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1042  | 
|
| 63654 | 1043  | 
lemma card_cartesian_product: "card (A \<times> B) = card A * card B"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1044  | 
by (cases "finite A \<and> finite B")  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1045  | 
(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
 | 
1046  | 
|
| 63654 | 1047  | 
lemma card_cartesian_product_singleton:  "card ({x} \<times> A) = card A"
 | 
1048  | 
by (simp add: card_cartesian_product)  | 
|
| 
54744
 
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  | 
|
| 60758 | 1051  | 
subsection \<open>Generalized product over a set\<close>  | 
| 
54744
 
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  | 
context comm_monoid_mult  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1054  | 
begin  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1055  | 
|
| 64272 | 1056  | 
sublocale prod: comm_monoid_set times 1  | 
1057  | 
defines prod = prod.F ..  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1058  | 
|
| 64272 | 1059  | 
abbreviation Prod ("\<Prod>_" [1000] 999)
 | 
1060  | 
where "\<Prod>A \<equiv> prod (\<lambda>x. x) A"  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1061  | 
|
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1062  | 
end  | 
| 
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1063  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
1064  | 
syntax (ASCII)  | 
| 
67268
 
bdf25939a550
new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
 
paulson <lp15@cam.ac.uk> 
parents: 
66936 
diff
changeset
 | 
1065  | 
  "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(4PROD (_/:_)./ _)" [0, 51, 10] 10)
 | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
1066  | 
syntax  | 
| 
67268
 
bdf25939a550
new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
 
paulson <lp15@cam.ac.uk> 
parents: 
66936 
diff
changeset
 | 
1067  | 
  "_prod" :: "pttrn => 'a set => 'b => 'b::comm_monoid_mult"  ("(2\<Prod>(_/\<in>_)./ _)" [0, 51, 10] 10)
 | 
| 61799 | 1068  | 
translations \<comment> \<open>Beware of argument permutation!\<close>  | 
| 64272 | 1069  | 
"\<Prod>i\<in>A. b" == "CONST prod (\<lambda>i. b) A"  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1070  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
1071  | 
text \<open>Instead of @{term"\<Prod>x\<in>{x. P}. e"} we introduce the shorter \<open>\<Prod>x|P. e\<close>.\<close>
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1072  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61952 
diff
changeset
 | 
1073  | 
syntax (ASCII)  | 
| 64272 | 1074  | 
  "_qprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(4PROD _ |/ _./ _)" [0, 0, 10] 10)
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1075  | 
syntax  | 
| 64272 | 1076  | 
  "_qprod" :: "pttrn \<Rightarrow> bool \<Rightarrow> 'a \<Rightarrow> 'a"  ("(2\<Prod>_ | (_)./ _)" [0, 0, 10] 10)
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1077  | 
translations  | 
| 64272 | 1078  | 
  "\<Prod>x|P. t" => "CONST prod (\<lambda>x. t) {x. P}"
 | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1079  | 
|
| 59010 | 1080  | 
context comm_monoid_mult  | 
1081  | 
begin  | 
|
1082  | 
||
| 64272 | 1083  | 
lemma prod_dvd_prod: "(\<And>a. a \<in> A \<Longrightarrow> f a dvd g a) \<Longrightarrow> prod f A dvd prod g A"  | 
| 59010 | 1084  | 
proof (induct A rule: infinite_finite_induct)  | 
| 63654 | 1085  | 
case infinite  | 
1086  | 
then show ?case by (auto intro: dvdI)  | 
|
1087  | 
next  | 
|
1088  | 
case empty  | 
|
1089  | 
then show ?case by (auto intro: dvdI)  | 
|
| 59010 | 1090  | 
next  | 
| 63654 | 1091  | 
case (insert a A)  | 
| 64272 | 1092  | 
then have "f a dvd g a" and "prod f A dvd prod g A"  | 
| 63654 | 1093  | 
by simp_all  | 
| 64272 | 1094  | 
then obtain r s where "g a = f a * r" and "prod g A = prod f A * s"  | 
| 63654 | 1095  | 
by (auto elim!: dvdE)  | 
| 64272 | 1096  | 
then have "g a * prod g A = f a * prod f A * (r * s)"  | 
| 63654 | 1097  | 
by (simp add: ac_simps)  | 
1098  | 
with insert.hyps show ?case  | 
|
1099  | 
by (auto intro: dvdI)  | 
|
| 59010 | 1100  | 
qed  | 
1101  | 
||
| 64272 | 1102  | 
lemma prod_dvd_prod_subset: "finite B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> prod f A dvd prod f B"  | 
1103  | 
by (auto simp add: prod.subset_diff ac_simps intro: dvdI)  | 
|
| 59010 | 1104  | 
|
1105  | 
end  | 
|
1106  | 
||
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1107  | 
|
| 60758 | 1108  | 
subsubsection \<open>Properties in more restricted classes of structures\<close>  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1109  | 
|
| 65687 | 1110  | 
context linordered_nonzero_semiring  | 
1111  | 
begin  | 
|
1112  | 
||
1113  | 
lemma prod_ge_1: "(\<And>x. x \<in> A \<Longrightarrow> 1 \<le> f x) \<Longrightarrow> 1 \<le> prod f A"  | 
|
1114  | 
proof (induct A rule: infinite_finite_induct)  | 
|
1115  | 
case infinite  | 
|
1116  | 
then show ?case by simp  | 
|
1117  | 
next  | 
|
1118  | 
case empty  | 
|
1119  | 
then show ?case by simp  | 
|
1120  | 
next  | 
|
1121  | 
case (insert x F)  | 
|
1122  | 
have "1 * 1 \<le> f x * prod f F"  | 
|
1123  | 
by (rule mult_mono') (use insert in auto)  | 
|
1124  | 
with insert show ?case by simp  | 
|
1125  | 
qed  | 
|
1126  | 
||
1127  | 
lemma prod_le_1:  | 
|
1128  | 
fixes f :: "'b \<Rightarrow> 'a"  | 
|
1129  | 
assumes "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x \<and> f x \<le> 1"  | 
|
1130  | 
shows "prod f A \<le> 1"  | 
|
1131  | 
using assms  | 
|
1132  | 
proof (induct A rule: infinite_finite_induct)  | 
|
1133  | 
case infinite  | 
|
1134  | 
then show ?case by simp  | 
|
1135  | 
next  | 
|
1136  | 
case empty  | 
|
1137  | 
then show ?case by simp  | 
|
1138  | 
next  | 
|
1139  | 
case (insert x F)  | 
|
1140  | 
then show ?case by (force simp: mult.commute intro: dest: mult_le_one)  | 
|
1141  | 
qed  | 
|
1142  | 
||
1143  | 
end  | 
|
1144  | 
||
| 59010 | 1145  | 
context comm_semiring_1  | 
1146  | 
begin  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1147  | 
|
| 64272 | 1148  | 
lemma dvd_prod_eqI [intro]:  | 
| 59010 | 1149  | 
assumes "finite A" and "a \<in> A" and "b = f a"  | 
| 64272 | 1150  | 
shows "b dvd prod f A"  | 
| 59010 | 1151  | 
proof -  | 
| 64272 | 1152  | 
  from \<open>finite A\<close> have "prod f (insert a (A - {a})) = f a * prod f (A - {a})"
 | 
1153  | 
by (intro prod.insert) auto  | 
|
| 63654 | 1154  | 
  also from \<open>a \<in> A\<close> have "insert a (A - {a}) = A"
 | 
1155  | 
by blast  | 
|
| 64272 | 1156  | 
  finally have "prod f A = f a * prod f (A - {a})" .
 | 
| 63654 | 1157  | 
with \<open>b = f a\<close> show ?thesis  | 
1158  | 
by simp  | 
|
| 59010 | 1159  | 
qed  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1160  | 
|
| 64272 | 1161  | 
lemma dvd_prodI [intro]: "finite A \<Longrightarrow> a \<in> A \<Longrightarrow> f a dvd prod f A"  | 
| 63654 | 1162  | 
by auto  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1163  | 
|
| 64272 | 1164  | 
lemma prod_zero:  | 
| 59010 | 1165  | 
assumes "finite A" and "\<exists>a\<in>A. f a = 0"  | 
| 64272 | 1166  | 
shows "prod f A = 0"  | 
| 63654 | 1167  | 
using assms  | 
1168  | 
proof (induct A)  | 
|
1169  | 
case empty  | 
|
1170  | 
then show ?case by simp  | 
|
| 59010 | 1171  | 
next  | 
1172  | 
case (insert a A)  | 
|
1173  | 
then have "f a = 0 \<or> (\<exists>a\<in>A. f a = 0)" by simp  | 
|
| 64272 | 1174  | 
then have "f a * prod f A = 0" by rule (simp_all add: insert)  | 
| 59010 | 1175  | 
with insert show ?case by simp  | 
1176  | 
qed  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1177  | 
|
| 64272 | 1178  | 
lemma prod_dvd_prod_subset2:  | 
| 59010 | 1179  | 
assumes "finite B" and "A \<subseteq> B" and "\<And>a. a \<in> A \<Longrightarrow> f a dvd g a"  | 
| 64272 | 1180  | 
shows "prod f A dvd prod g B"  | 
| 59010 | 1181  | 
proof -  | 
| 64272 | 1182  | 
from assms have "prod f A dvd prod g A"  | 
1183  | 
by (auto intro: prod_dvd_prod)  | 
|
1184  | 
moreover from assms have "prod g A dvd prod g B"  | 
|
1185  | 
by (auto intro: prod_dvd_prod_subset)  | 
|
| 59010 | 1186  | 
ultimately show ?thesis by (rule dvd_trans)  | 
1187  | 
qed  | 
|
1188  | 
||
1189  | 
end  | 
|
1190  | 
||
| 64272 | 1191  | 
lemma (in semidom) prod_zero_iff [simp]:  | 
| 63924 | 1192  | 
fixes f :: "'b \<Rightarrow> 'a"  | 
| 59010 | 1193  | 
assumes "finite A"  | 
| 64272 | 1194  | 
shows "prod f A = 0 \<longleftrightarrow> (\<exists>a\<in>A. f a = 0)"  | 
| 59010 | 1195  | 
using assms by (induct A) (auto simp: no_zero_divisors)  | 
1196  | 
||
| 64272 | 1197  | 
lemma (in semidom_divide) prod_diff1:  | 
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1198  | 
assumes "finite A" and "f a \<noteq> 0"  | 
| 64272 | 1199  | 
  shows "prod f (A - {a}) = (if a \<in> A then prod f A div f a else prod f A)"
 | 
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1200  | 
proof (cases "a \<notin> A")  | 
| 63654 | 1201  | 
case True  | 
1202  | 
then show ?thesis by simp  | 
|
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1203  | 
next  | 
| 63654 | 1204  | 
case False  | 
1205  | 
with assms show ?thesis  | 
|
1206  | 
proof induct  | 
|
1207  | 
case empty  | 
|
1208  | 
then show ?case by simp  | 
|
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1209  | 
next  | 
| 
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1210  | 
case (insert b B)  | 
| 
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1211  | 
then show ?case  | 
| 
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1212  | 
proof (cases "a = b")  | 
| 63654 | 1213  | 
case True  | 
1214  | 
with insert show ?thesis by simp  | 
|
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1215  | 
next  | 
| 63654 | 1216  | 
case False  | 
1217  | 
with insert have "a \<in> B" by simp  | 
|
| 63040 | 1218  | 
      define C where "C = B - {a}"
 | 
| 63654 | 1219  | 
with \<open>finite B\<close> \<open>a \<in> B\<close> have "B = insert a C" "finite C" "a \<notin> C"  | 
1220  | 
by auto  | 
|
1221  | 
with insert show ?thesis  | 
|
1222  | 
by (auto simp add: insert_commute ac_simps)  | 
|
| 
60353
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1223  | 
qed  | 
| 
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1224  | 
qed  | 
| 
 
838025c6e278
implicit partial divison operation in integral domains
 
haftmann 
parents: 
59867 
diff
changeset
 | 
1225  | 
qed  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1226  | 
|
| 64267 | 1227  | 
lemma sum_zero_power [simp]: "(\<Sum>i\<in>A. c i * 0^i) = (if finite A \<and> 0 \<in> A then c 0 else 0)"  | 
| 63654 | 1228  | 
for c :: "nat \<Rightarrow> 'a::division_ring"  | 
1229  | 
by (induct A rule: infinite_finite_induct) auto  | 
|
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1230  | 
|
| 64267 | 1231  | 
lemma sum_zero_power' [simp]:  | 
| 63654 | 1232  | 
"(\<Sum>i\<in>A. c i * 0^i / d i) = (if finite A \<and> 0 \<in> A then c 0 / d 0 else 0)"  | 
1233  | 
for c :: "nat \<Rightarrow> 'a::field"  | 
|
| 64267 | 1234  | 
using sum_zero_power [of "\<lambda>i. c i / d i" A] by auto  | 
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1235  | 
|
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1236  | 
lemma (in field) prod_inversef: "prod (inverse \<circ> f) A = inverse (prod f A)"  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1237  | 
proof (cases "finite A")  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1238  | 
case True  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1239  | 
then show ?thesis  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1240  | 
by (induct A rule: finite_induct) simp_all  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1241  | 
next  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1242  | 
case False  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1243  | 
then show ?thesis  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1244  | 
by auto  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1245  | 
qed  | 
| 59010 | 1246  | 
|
| 
65680
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1247  | 
lemma (in field) prod_dividef: "(\<Prod>x\<in>A. f x / g x) = prod f A / prod g A"  | 
| 
 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 
paulson <lp15@cam.ac.uk> 
parents: 
64979 
diff
changeset
 | 
1248  | 
using prod_inversef [of g A] by (simp add: divide_inverse prod.distrib)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1249  | 
|
| 64272 | 1250  | 
lemma prod_Un:  | 
| 59010 | 1251  | 
fixes f :: "'b \<Rightarrow> 'a :: field"  | 
1252  | 
assumes "finite A" and "finite B"  | 
|
| 63654 | 1253  | 
and "\<forall>x\<in>A \<inter> B. f x \<noteq> 0"  | 
| 64272 | 1254  | 
shows "prod f (A \<union> B) = prod f A * prod f B / prod f (A \<inter> B)"  | 
| 59010 | 1255  | 
proof -  | 
| 64272 | 1256  | 
from assms have "prod f A * prod f B = prod f (A \<union> B) * prod f (A \<inter> B)"  | 
1257  | 
by (simp add: prod.union_inter [symmetric, of A B])  | 
|
| 63654 | 1258  | 
with assms show ?thesis  | 
1259  | 
by simp  | 
|
| 59010 | 1260  | 
qed  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1261  | 
|
| 64272 | 1262  | 
lemma (in linordered_semidom) prod_nonneg: "(\<forall>a\<in>A. 0 \<le> f a) \<Longrightarrow> 0 \<le> prod f A"  | 
| 59010 | 1263  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
1264  | 
||
| 64272 | 1265  | 
lemma (in linordered_semidom) prod_pos: "(\<forall>a\<in>A. 0 < f a) \<Longrightarrow> 0 < prod f A"  | 
| 59010 | 1266  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
1267  | 
||
| 64272 | 1268  | 
lemma (in linordered_semidom) prod_mono:  | 
1269  | 
"\<forall>i\<in>A. 0 \<le> f i \<and> f i \<le> g i \<Longrightarrow> prod f A \<le> prod g A"  | 
|
1270  | 
by (induct A rule: infinite_finite_induct) (auto intro!: prod_nonneg mult_mono)  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1271  | 
|
| 64272 | 1272  | 
lemma (in linordered_semidom) prod_mono_strict:  | 
| 63654 | 1273  | 
  assumes "finite A" "\<forall>i\<in>A. 0 \<le> f i \<and> f i < g i" "A \<noteq> {}"
 | 
| 64272 | 1274  | 
shows "prod f A < prod g A"  | 
| 63654 | 1275  | 
using assms  | 
1276  | 
proof (induct A rule: finite_induct)  | 
|
1277  | 
case empty  | 
|
1278  | 
then show ?case by simp  | 
|
1279  | 
next  | 
|
1280  | 
case insert  | 
|
| 64272 | 1281  | 
then show ?case by (force intro: mult_strict_mono' prod_nonneg)  | 
| 63654 | 1282  | 
qed  | 
| 
60974
 
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
 
paulson <lp15@cam.ac.uk> 
parents: 
60758 
diff
changeset
 | 
1283  | 
|
| 64272 | 1284  | 
lemma (in linordered_field) abs_prod: "\<bar>prod f A\<bar> = (\<Prod>x\<in>A. \<bar>f x\<bar>)"  | 
| 59010 | 1285  | 
by (induct A rule: infinite_finite_induct) (simp_all add: abs_mult)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1286  | 
|
| 64272 | 1287  | 
lemma prod_eq_1_iff [simp]: "finite A \<Longrightarrow> prod f A = 1 \<longleftrightarrow> (\<forall>a\<in>A. f a = 1)"  | 
| 63654 | 1288  | 
for f :: "'a \<Rightarrow> nat"  | 
| 59010 | 1289  | 
by (induct A rule: finite_induct) simp_all  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1290  | 
|
| 64272 | 1291  | 
lemma prod_pos_nat_iff [simp]: "finite A \<Longrightarrow> prod f A > 0 \<longleftrightarrow> (\<forall>a\<in>A. f a > 0)"  | 
| 63654 | 1292  | 
for f :: "'a \<Rightarrow> nat"  | 
| 64272 | 1293  | 
using prod_zero_iff by (simp del: neq0_conv add: zero_less_iff_neq_zero)  | 
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1294  | 
|
| 64272 | 1295  | 
lemma prod_constant: "(\<Prod>x\<in> A. y) = y ^ card A"  | 
| 63654 | 1296  | 
for y :: "'a::comm_monoid_mult"  | 
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1297  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
| 
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1298  | 
|
| 64272 | 1299  | 
lemma prod_power_distrib: "prod f A ^ n = prod (\<lambda>x. (f x) ^ n) A"  | 
| 63654 | 1300  | 
for f :: "'a \<Rightarrow> 'b::comm_semiring_1"  | 
1301  | 
by (induct A rule: infinite_finite_induct) (auto simp add: power_mult_distrib)  | 
|
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1302  | 
|
| 64267 | 1303  | 
lemma power_sum: "c ^ (\<Sum>a\<in>A. f a) = (\<Prod>a\<in>A. c ^ f a)"  | 
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1304  | 
by (induct A rule: infinite_finite_induct) (simp_all add: power_add)  | 
| 
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1305  | 
|
| 64272 | 1306  | 
lemma prod_gen_delta:  | 
| 63654 | 1307  | 
fixes b :: "'b \<Rightarrow> 'a::comm_monoid_mult"  | 
1308  | 
assumes fin: "finite S"  | 
|
| 64272 | 1309  | 
shows "prod (\<lambda>k. if k = a then b k else c) S =  | 
| 63654 | 1310  | 
(if a \<in> S then b a * c ^ (card S - 1) else c ^ card S)"  | 
1311  | 
proof -  | 
|
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1312  | 
let ?f = "(\<lambda>k. if k=a then b k else c)"  | 
| 63654 | 1313  | 
show ?thesis  | 
1314  | 
proof (cases "a \<in> S")  | 
|
1315  | 
case False  | 
|
1316  | 
then have "\<forall> k\<in> S. ?f k = c" by simp  | 
|
| 64272 | 1317  | 
with False show ?thesis by (simp add: prod_constant)  | 
| 63654 | 1318  | 
next  | 
1319  | 
case True  | 
|
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1320  | 
    let ?A = "S - {a}"
 | 
| 
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1321  | 
    let ?B = "{a}"
 | 
| 63654 | 1322  | 
from True have eq: "S = ?A \<union> ?B" by blast  | 
1323  | 
    have disjoint: "?A \<inter> ?B = {}" by simp
 | 
|
1324  | 
from fin have fin': "finite ?A" "finite ?B" by auto  | 
|
| 64272 | 1325  | 
have f_A0: "prod ?f ?A = prod (\<lambda>i. c) ?A"  | 
1326  | 
by (rule prod.cong) auto  | 
|
| 63654 | 1327  | 
from fin True have card_A: "card ?A = card S - 1" by auto  | 
| 64272 | 1328  | 
have f_A1: "prod ?f ?A = c ^ card ?A"  | 
1329  | 
unfolding f_A0 by (rule prod_constant)  | 
|
1330  | 
have "prod ?f ?A * prod ?f ?B = prod ?f S"  | 
|
1331  | 
using prod.union_disjoint[OF fin' disjoint, of ?f, unfolded eq[symmetric]]  | 
|
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1332  | 
by simp  | 
| 63654 | 1333  | 
with True card_A show ?thesis  | 
| 64272 | 1334  | 
by (simp add: f_A1 field_simps cong add: prod.cong cong del: if_weak_cong)  | 
| 63654 | 1335  | 
qed  | 
| 
62481
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1336  | 
qed  | 
| 
 
b5d8e57826df
tuned bootstrap order to provide type classes in a more sensible order
 
haftmann 
parents: 
62378 
diff
changeset
 | 
1337  | 
|
| 64267 | 1338  | 
lemma sum_image_le:  | 
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1339  | 
fixes g :: "'a \<Rightarrow> 'b::ordered_ab_group_add"  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1340  | 
assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> g(f i)"  | 
| 64267 | 1341  | 
shows "sum g (f ` I) \<le> sum (g \<circ> f) I"  | 
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1342  | 
using assms  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1343  | 
proof induction  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1344  | 
case empty  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1345  | 
then show ?case by auto  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1346  | 
next  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1347  | 
case (insert x F) then  | 
| 64267 | 1348  | 
have "sum g (f ` insert x F) = sum g (insert (f x) (f ` F))" by simp  | 
1349  | 
also have "\<dots> \<le> g (f x) + sum g (f ` F)"  | 
|
1350  | 
by (simp add: insert sum.insert_if)  | 
|
1351  | 
also have "\<dots> \<le> sum (g \<circ> f) (insert x F)"  | 
|
| 
63952
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1352  | 
using insert by auto  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1353  | 
finally show ?case .  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1354  | 
qed  | 
| 
 
354808e9f44b
new material connected with HOL Light measure theory, plus more rationalisation
 
paulson <lp15@cam.ac.uk> 
parents: 
63938 
diff
changeset
 | 
1355  | 
|
| 
54744
 
1e7f2d296e19
more algebraic terminology for theories about big operators
 
haftmann 
parents:  
diff
changeset
 | 
1356  | 
end  |