author | wenzelm |
Fri, 19 Jul 2024 11:29:05 +0200 | |
changeset 80589 | 7849b6370425 |
parent 80095 | 0f9cd1a5edbe |
child 80768 | c7723cc15de8 |
permissions | -rw-r--r-- |
58197 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
2 |
||
58881 | 3 |
section \<open>Big sum and product over function bodies\<close> |
58197 | 4 |
|
5 |
theory Groups_Big_Fun |
|
6 |
imports |
|
7 |
Main |
|
8 |
begin |
|
9 |
||
10 |
subsection \<open>Abstract product\<close> |
|
11 |
||
12 |
locale comm_monoid_fun = comm_monoid |
|
13 |
begin |
|
14 |
||
15 |
definition G :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
16 |
where |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
17 |
expand_set: "G g = comm_monoid_set.F f \<^bold>1 g {a. g a \<noteq> \<^bold>1}" |
58197 | 18 |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
19 |
interpretation F: comm_monoid_set f "\<^bold>1" |
58197 | 20 |
.. |
21 |
||
22 |
lemma expand_superset: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
23 |
assumes "finite A" and "{a. g a \<noteq> \<^bold>1} \<subseteq> A" |
58197 | 24 |
shows "G g = F.F g A" |
80095 | 25 |
using F.mono_neutral_right assms expand_set by fastforce |
58197 | 26 |
|
27 |
lemma conditionalize: |
|
28 |
assumes "finite A" |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
29 |
shows "F.F g A = G (\<lambda>a. if a \<in> A then g a else \<^bold>1)" |
58197 | 30 |
using assms |
80095 | 31 |
by (smt (verit, ccfv_threshold) Diff_iff F.mono_neutral_cong_right expand_set mem_Collect_eq subsetI) |
32 |
||
58197 | 33 |
|
34 |
lemma neutral [simp]: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
35 |
"G (\<lambda>a. \<^bold>1) = \<^bold>1" |
58197 | 36 |
by (simp add: expand_set) |
37 |
||
38 |
lemma update [simp]: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
39 |
assumes "finite {a. g a \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
40 |
assumes "g a = \<^bold>1" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
41 |
shows "G (g(a := b)) = b \<^bold>* G g" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
42 |
proof (cases "b = \<^bold>1") |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
43 |
case True with \<open>g a = \<^bold>1\<close> show ?thesis |
58197 | 44 |
by (simp add: expand_set) (rule F.cong, auto) |
45 |
next |
|
46 |
case False |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
47 |
moreover have "{a'. a' \<noteq> a \<longrightarrow> g a' \<noteq> \<^bold>1} = insert a {a. g a \<noteq> \<^bold>1}" |
58197 | 48 |
by auto |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
49 |
moreover from \<open>g a = \<^bold>1\<close> have "a \<notin> {a. g a \<noteq> \<^bold>1}" |
58197 | 50 |
by simp |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
51 |
moreover have "F.F (\<lambda>a'. if a' = a then b else g a') {a. g a \<noteq> \<^bold>1} = F.F g {a. g a \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
52 |
by (rule F.cong) (auto simp add: \<open>g a = \<^bold>1\<close>) |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
53 |
ultimately show ?thesis using \<open>finite {a. g a \<noteq> \<^bold>1}\<close> by (simp add: expand_set) |
58197 | 54 |
qed |
55 |
||
56 |
lemma infinite [simp]: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
57 |
"\<not> finite {a. g a \<noteq> \<^bold>1} \<Longrightarrow> G g = \<^bold>1" |
58197 | 58 |
by (simp add: expand_set) |
59 |
||
69164 | 60 |
lemma cong [cong]: |
58197 | 61 |
assumes "\<And>a. g a = h a" |
62 |
shows "G g = G h" |
|
63 |
using assms by (simp add: expand_set) |
|
64 |
||
65 |
lemma not_neutral_obtains_not_neutral: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
66 |
assumes "G g \<noteq> \<^bold>1" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
67 |
obtains a where "g a \<noteq> \<^bold>1" |
58197 | 68 |
using assms by (auto elim: F.not_neutral_contains_not_neutral simp add: expand_set) |
69 |
||
70 |
lemma reindex_cong: |
|
71 |
assumes "bij l" |
|
72 |
assumes "g \<circ> l = h" |
|
73 |
shows "G g = G h" |
|
74 |
proof - |
|
75 |
from assms have unfold: "h = g \<circ> l" by simp |
|
60500 | 76 |
from \<open>bij l\<close> have "inj l" by (rule bij_is_inj) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
77 |
then have "inj_on l {a. h a \<noteq> \<^bold>1}" by (rule subset_inj_on) simp |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
78 |
moreover from \<open>bij l\<close> have "{a. g a \<noteq> \<^bold>1} = l ` {a. h a \<noteq> \<^bold>1}" |
58197 | 79 |
by (auto simp add: image_Collect unfold elim: bij_pointE) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
80 |
moreover have "\<And>x. x \<in> {a. h a \<noteq> \<^bold>1} \<Longrightarrow> g (l x) = h x" |
58197 | 81 |
by (simp add: unfold) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
82 |
ultimately have "F.F g {a. g a \<noteq> \<^bold>1} = F.F h {a. h a \<noteq> \<^bold>1}" |
58197 | 83 |
by (rule F.reindex_cong) |
84 |
then show ?thesis by (simp add: expand_set) |
|
85 |
qed |
|
86 |
||
87 |
lemma distrib: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
88 |
assumes "finite {a. g a \<noteq> \<^bold>1}" and "finite {a. h a \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
89 |
shows "G (\<lambda>a. g a \<^bold>* h a) = G g \<^bold>* G h" |
58197 | 90 |
proof - |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
91 |
from assms have "finite ({a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1})" by simp |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
92 |
moreover have "{a. g a \<^bold>* h a \<noteq> \<^bold>1} \<subseteq> {a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}" |
58197 | 93 |
by auto (drule sym, simp) |
94 |
ultimately show ?thesis |
|
95 |
using assms |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
96 |
by (simp add: expand_superset [of "{a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}"] F.distrib) |
58197 | 97 |
qed |
98 |
||
66804
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
haftmann
parents:
64272
diff
changeset
|
99 |
lemma swap: |
58197 | 100 |
assumes "finite C" |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
101 |
assumes subset: "{a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C") |
58197 | 102 |
shows "G (\<lambda>a. G (g a)) = G (\<lambda>b. G (\<lambda>a. g a b))" |
103 |
proof - |
|
60500 | 104 |
from \<open>finite C\<close> subset |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
105 |
have "finite ({a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1})" |
58197 | 106 |
by (rule rev_finite_subset) |
107 |
then have fins: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
108 |
"finite {b. \<exists>a. g a b \<noteq> \<^bold>1}" "finite {a. \<exists>b. g a b \<noteq> \<^bold>1}" |
58197 | 109 |
by (auto simp add: finite_cartesian_product_iff) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
110 |
have subsets: "\<And>a. {b. g a b \<noteq> \<^bold>1} \<subseteq> {b. \<exists>a. g a b \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
111 |
"\<And>b. {a. g a b \<noteq> \<^bold>1} \<subseteq> {a. \<exists>b. g a b \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
112 |
"{a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {a. \<exists>b. g a b \<noteq> \<^bold>1}" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
113 |
"{a. F.F (\<lambda>aa. g aa a) {a. \<exists>b. g a b \<noteq> \<^bold>1} \<noteq> \<^bold>1} \<subseteq> {b. \<exists>a. g a b \<noteq> \<^bold>1}" |
58197 | 114 |
by (auto elim: F.not_neutral_contains_not_neutral) |
66804
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
haftmann
parents:
64272
diff
changeset
|
115 |
from F.swap have |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
116 |
"F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1} = |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
117 |
F.F (\<lambda>b. F.F (\<lambda>a. g a b) {a. \<exists>b. g a b \<noteq> \<^bold>1}) {b. \<exists>a. g a b \<noteq> \<^bold>1}" . |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
118 |
with subsets fins have "G (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) = |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
119 |
G (\<lambda>b. F.F (\<lambda>a. g a b) {a. \<exists>b. g a b \<noteq> \<^bold>1})" |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
120 |
by (auto simp add: expand_superset [of "{b. \<exists>a. g a b \<noteq> \<^bold>1}"] |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
121 |
expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"]) |
58197 | 122 |
with subsets fins show ?thesis |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
123 |
by (auto simp add: expand_superset [of "{b. \<exists>a. g a b \<noteq> \<^bold>1}"] |
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
124 |
expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"]) |
58197 | 125 |
qed |
126 |
||
127 |
lemma cartesian_product: |
|
128 |
assumes "finite C" |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
129 |
assumes subset: "{a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1} \<subseteq> C" (is "?A \<times> ?B \<subseteq> C") |
58197 | 130 |
shows "G (\<lambda>a. G (g a)) = G (\<lambda>(a, b). g a b)" |
131 |
proof - |
|
60500 | 132 |
from subset \<open>finite C\<close> have fin_prod: "finite (?A \<times> ?B)" |
58197 | 133 |
by (rule finite_subset) |
134 |
from fin_prod have "finite ?A" and "finite ?B" |
|
135 |
by (auto simp add: finite_cartesian_product_iff) |
|
136 |
have *: "G (\<lambda>a. G (g a)) = |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
137 |
(F.F (\<lambda>a. F.F (g a) {b. \<exists>a. g a b \<noteq> \<^bold>1}) {a. \<exists>b. g a b \<noteq> \<^bold>1})" |
80095 | 138 |
using \<open>finite ?A\<close> \<open>finite ?B\<close> expand_superset |
139 |
by (smt (verit, del_insts) Collect_mono local.cong not_neutral_obtains_not_neutral) |
|
140 |
have **: "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> ?A \<times> ?B" |
|
58197 | 141 |
by auto |
142 |
show ?thesis |
|
80095 | 143 |
using \<open>finite C\<close> expand_superset |
144 |
using "*" ** F.cartesian_product fin_prod by force |
|
58197 | 145 |
qed |
146 |
||
147 |
lemma cartesian_product2: |
|
148 |
assumes fin: "finite D" |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
149 |
assumes subset: "{(a, b). \<exists>c. g a b c \<noteq> \<^bold>1} \<times> {c. \<exists>a b. g a b c \<noteq> \<^bold>1} \<subseteq> D" (is "?AB \<times> ?C \<subseteq> D") |
58197 | 150 |
shows "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>(a, b, c). g a b c)" |
151 |
proof - |
|
152 |
have bij: "bij (\<lambda>(a, b, c). ((a, b), c))" |
|
153 |
by (auto intro!: bijI injI simp add: image_def) |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
154 |
have "{p. \<exists>c. g (fst p) (snd p) c \<noteq> \<^bold>1} \<times> {c. \<exists>p. g (fst p) (snd p) c \<noteq> \<^bold>1} \<subseteq> D" |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61378
diff
changeset
|
155 |
by auto (insert subset, blast) |
58197 | 156 |
with fin have "G (\<lambda>p. G (g (fst p) (snd p))) = G (\<lambda>(p, c). g (fst p) (snd p) c)" |
157 |
by (rule cartesian_product) |
|
158 |
then have "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>((a, b), c). g a b c)" |
|
159 |
by (auto simp add: split_def) |
|
160 |
also have "G (\<lambda>((a, b), c). g a b c) = G (\<lambda>(a, b, c). g a b c)" |
|
161 |
using bij by (rule reindex_cong [of "\<lambda>(a, b, c). ((a, b), c)"]) (simp add: fun_eq_iff) |
|
162 |
finally show ?thesis . |
|
163 |
qed |
|
164 |
||
165 |
lemma delta [simp]: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
166 |
"G (\<lambda>b. if b = a then g b else \<^bold>1) = g a" |
58197 | 167 |
proof - |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
168 |
have "{b. (if b = a then g b else \<^bold>1) \<noteq> \<^bold>1} \<subseteq> {a}" by auto |
58197 | 169 |
then show ?thesis by (simp add: expand_superset [of "{a}"]) |
170 |
qed |
|
171 |
||
172 |
lemma delta' [simp]: |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
173 |
"G (\<lambda>b. if a = b then g b else \<^bold>1) = g a" |
58197 | 174 |
proof - |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
175 |
have "(\<lambda>b. if a = b then g b else \<^bold>1) = (\<lambda>b. if b = a then g b else \<^bold>1)" |
58197 | 176 |
by (simp add: fun_eq_iff) |
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
61955
diff
changeset
|
177 |
then have "G (\<lambda>b. if a = b then g b else \<^bold>1) = G (\<lambda>b. if b = a then g b else \<^bold>1)" |
69164 | 178 |
by (simp cong del: cong) |
58197 | 179 |
then show ?thesis by simp |
180 |
qed |
|
181 |
||
182 |
end |
|
183 |
||
184 |
||
185 |
subsection \<open>Concrete sum\<close> |
|
186 |
||
187 |
context comm_monoid_add |
|
188 |
begin |
|
189 |
||
61776 | 190 |
sublocale Sum_any: comm_monoid_fun plus 0 |
67764 | 191 |
rewrites "comm_monoid_set.F plus 0 = sum" |
63433 | 192 |
defines Sum_any = Sum_any.G |
58197 | 193 |
proof - |
194 |
show "comm_monoid_fun plus 0" .. |
|
61605 | 195 |
then interpret Sum_any: comm_monoid_fun plus 0 . |
64267 | 196 |
from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym) |
58197 | 197 |
qed |
198 |
||
199 |
end |
|
200 |
||
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
201 |
syntax (ASCII) |
58197 | 202 |
"_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10) |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
203 |
syntax |
58197 | 204 |
"_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3\<Sum>_. _)" [0, 10] 10) |
205 |
translations |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
206 |
"\<Sum>a. b" \<rightleftharpoons> "CONST Sum_any (\<lambda>a. b)" |
58197 | 207 |
|
208 |
lemma Sum_any_left_distrib: |
|
209 |
fixes r :: "'a :: semiring_0" |
|
210 |
assumes "finite {a. g a \<noteq> 0}" |
|
211 |
shows "Sum_any g * r = (\<Sum>n. g n * r)" |
|
80095 | 212 |
by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_left sum_distrib_right) |
58197 | 213 |
|
214 |
lemma Sum_any_right_distrib: |
|
215 |
fixes r :: "'a :: semiring_0" |
|
216 |
assumes "finite {a. g a \<noteq> 0}" |
|
217 |
shows "r * Sum_any g = (\<Sum>n. r * g n)" |
|
80095 | 218 |
by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_right sum_distrib_left) |
58197 | 219 |
|
220 |
lemma Sum_any_product: |
|
221 |
fixes f g :: "'b \<Rightarrow> 'a::semiring_0" |
|
222 |
assumes "finite {a. f a \<noteq> 0}" and "finite {b. g b \<noteq> 0}" |
|
223 |
shows "Sum_any f * Sum_any g = (\<Sum>a. \<Sum>b. f a * g b)" |
|
224 |
proof - |
|
80095 | 225 |
have "\<forall>a. (\<Sum>b. a * g b) = a * Sum_any g" |
226 |
by (simp add: Sum_any_right_distrib assms(2)) |
|
227 |
then show ?thesis |
|
228 |
by (simp add: Sum_any_left_distrib assms(1)) |
|
58197 | 229 |
qed |
230 |
||
58437 | 231 |
lemma Sum_any_eq_zero_iff [simp]: |
232 |
fixes f :: "'a \<Rightarrow> nat" |
|
233 |
assumes "finite {a. f a \<noteq> 0}" |
|
234 |
shows "Sum_any f = 0 \<longleftrightarrow> f = (\<lambda>_. 0)" |
|
235 |
using assms by (simp add: Sum_any.expand_set fun_eq_iff) |
|
236 |
||
58197 | 237 |
|
238 |
subsection \<open>Concrete product\<close> |
|
239 |
||
240 |
context comm_monoid_mult |
|
241 |
begin |
|
242 |
||
61776 | 243 |
sublocale Prod_any: comm_monoid_fun times 1 |
67764 | 244 |
rewrites "comm_monoid_set.F times 1 = prod" |
63433 | 245 |
defines Prod_any = Prod_any.G |
58197 | 246 |
proof - |
247 |
show "comm_monoid_fun times 1" .. |
|
61605 | 248 |
then interpret Prod_any: comm_monoid_fun times 1 . |
64272 | 249 |
from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym) |
58197 | 250 |
qed |
251 |
||
252 |
end |
|
253 |
||
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
254 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
255 |
"_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) |
58197 | 256 |
syntax |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61776
diff
changeset
|
257 |
"_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" ("(3\<Prod>_. _)" [0, 10] 10) |
58197 | 258 |
translations |
259 |
"\<Prod>a. b" == "CONST Prod_any (\<lambda>a. b)" |
|
260 |
||
261 |
lemma Prod_any_zero: |
|
262 |
fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1" |
|
263 |
assumes "finite {a. f a \<noteq> 1}" |
|
264 |
assumes "f a = 0" |
|
265 |
shows "(\<Prod>a. f a) = 0" |
|
266 |
proof - |
|
60500 | 267 |
from \<open>f a = 0\<close> have "f a \<noteq> 1" by simp |
268 |
with \<open>f a = 0\<close> have "\<exists>a. f a \<noteq> 1 \<and> f a = 0" by blast |
|
269 |
with \<open>finite {a. f a \<noteq> 1}\<close> show ?thesis |
|
64272 | 270 |
by (simp add: Prod_any.expand_set prod_zero) |
58197 | 271 |
qed |
272 |
||
273 |
lemma Prod_any_not_zero: |
|
274 |
fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1" |
|
275 |
assumes "finite {a. f a \<noteq> 1}" |
|
276 |
assumes "(\<Prod>a. f a) \<noteq> 0" |
|
277 |
shows "f a \<noteq> 0" |
|
278 |
using assms Prod_any_zero [of f] by blast |
|
279 |
||
58437 | 280 |
lemma power_Sum_any: |
281 |
assumes "finite {a. f a \<noteq> 0}" |
|
282 |
shows "c ^ (\<Sum>a. f a) = (\<Prod>a. c ^ f a)" |
|
283 |
proof - |
|
284 |
have "{a. c ^ f a \<noteq> 1} \<subseteq> {a. f a \<noteq> 0}" |
|
285 |
by (auto intro: ccontr) |
|
286 |
with assms show ?thesis |
|
64267 | 287 |
by (simp add: Sum_any.expand_set Prod_any.expand_superset power_sum) |
58437 | 288 |
qed |
289 |
||
58197 | 290 |
end |