| author | paulson <lp15@cam.ac.uk> | 
| Tue, 02 Apr 2019 12:56:05 +0100 | |
| changeset 70027 | 94494b92d8d0 | 
| parent 69164 | 74f1b0f10b2b | 
| child 80095 | 0f9cd1a5edbe | 
| 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"  | 
25  | 
apply (simp add: expand_set)  | 
|
26  | 
apply (rule F.same_carrierI [of A])  | 
|
27  | 
apply (simp_all add: assms)  | 
|
28  | 
done  | 
|
29  | 
||
30  | 
lemma conditionalize:  | 
|
31  | 
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
 | 
32  | 
shows "F.F g A = G (\<lambda>a. if a \<in> A then g a else \<^bold>1)"  | 
| 58197 | 33  | 
using assms  | 
34  | 
apply (simp add: expand_set)  | 
|
35  | 
apply (rule F.same_carrierI [of A])  | 
|
36  | 
apply auto  | 
|
37  | 
done  | 
|
38  | 
||
39  | 
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
 | 
40  | 
"G (\<lambda>a. \<^bold>1) = \<^bold>1"  | 
| 58197 | 41  | 
by (simp add: expand_set)  | 
42  | 
||
43  | 
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
 | 
44  | 
  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
 | 
45  | 
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
 | 
46  | 
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
 | 
47  | 
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
 | 
48  | 
case True with \<open>g a = \<^bold>1\<close> show ?thesis  | 
| 58197 | 49  | 
by (simp add: expand_set) (rule F.cong, auto)  | 
50  | 
next  | 
|
51  | 
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
 | 
52  | 
  moreover have "{a'. a' \<noteq> a \<longrightarrow> g a' \<noteq> \<^bold>1} = insert a {a. g a \<noteq> \<^bold>1}"
 | 
| 58197 | 53  | 
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
 | 
54  | 
  moreover from \<open>g a = \<^bold>1\<close> have "a \<notin> {a. g a \<noteq> \<^bold>1}"
 | 
| 58197 | 55  | 
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
 | 
56  | 
  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
 | 
57  | 
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
 | 
58  | 
  ultimately show ?thesis using \<open>finite {a. g a \<noteq> \<^bold>1}\<close> by (simp add: expand_set)
 | 
| 58197 | 59  | 
qed  | 
60  | 
||
61  | 
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
 | 
62  | 
  "\<not> finite {a. g a \<noteq> \<^bold>1} \<Longrightarrow> G g = \<^bold>1"
 | 
| 58197 | 63  | 
by (simp add: expand_set)  | 
64  | 
||
| 69164 | 65  | 
lemma cong [cong]:  | 
| 58197 | 66  | 
assumes "\<And>a. g a = h a"  | 
67  | 
shows "G g = G h"  | 
|
68  | 
using assms by (simp add: expand_set)  | 
|
69  | 
||
70  | 
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
 | 
71  | 
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
 | 
72  | 
obtains a where "g a \<noteq> \<^bold>1"  | 
| 58197 | 73  | 
using assms by (auto elim: F.not_neutral_contains_not_neutral simp add: expand_set)  | 
74  | 
||
75  | 
lemma reindex_cong:  | 
|
76  | 
assumes "bij l"  | 
|
77  | 
assumes "g \<circ> l = h"  | 
|
78  | 
shows "G g = G h"  | 
|
79  | 
proof -  | 
|
80  | 
from assms have unfold: "h = g \<circ> l" by simp  | 
|
| 60500 | 81  | 
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
 | 
82  | 
  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
 | 
83  | 
  moreover from \<open>bij l\<close> have "{a. g a \<noteq> \<^bold>1} = l ` {a. h a \<noteq> \<^bold>1}"
 | 
| 58197 | 84  | 
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
 | 
85  | 
  moreover have "\<And>x. x \<in> {a. h a \<noteq> \<^bold>1} \<Longrightarrow> g (l x) = h x"
 | 
| 58197 | 86  | 
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
 | 
87  | 
  ultimately have "F.F g {a. g a \<noteq> \<^bold>1} = F.F h {a. h a \<noteq> \<^bold>1}"
 | 
| 58197 | 88  | 
by (rule F.reindex_cong)  | 
89  | 
then show ?thesis by (simp add: expand_set)  | 
|
90  | 
qed  | 
|
91  | 
||
92  | 
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
 | 
93  | 
  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
 | 
94  | 
shows "G (\<lambda>a. g a \<^bold>* h a) = G g \<^bold>* G h"  | 
| 58197 | 95  | 
proof -  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
61955 
diff
changeset
 | 
96  | 
  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
 | 
97  | 
  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 | 98  | 
by auto (drule sym, simp)  | 
99  | 
ultimately show ?thesis  | 
|
100  | 
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
 | 
101  | 
    by (simp add: expand_superset [of "{a. g a \<noteq> \<^bold>1} \<union> {a. h a \<noteq> \<^bold>1}"] F.distrib)
 | 
| 58197 | 102  | 
qed  | 
103  | 
||
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
64272 
diff
changeset
 | 
104  | 
lemma swap:  | 
| 58197 | 105  | 
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
 | 
106  | 
  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 | 107  | 
shows "G (\<lambda>a. G (g a)) = G (\<lambda>b. G (\<lambda>a. g a b))"  | 
108  | 
proof -  | 
|
| 60500 | 109  | 
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
 | 
110  | 
    have "finite ({a. \<exists>b. g a b \<noteq> \<^bold>1} \<times> {b. \<exists>a. g a b \<noteq> \<^bold>1})"
 | 
| 58197 | 111  | 
by (rule rev_finite_subset)  | 
112  | 
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
 | 
113  | 
    "finite {b. \<exists>a. g a b \<noteq> \<^bold>1}" "finite {a. \<exists>b. g a b \<noteq> \<^bold>1}"
 | 
| 58197 | 114  | 
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
 | 
115  | 
  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
 | 
116  | 
    "\<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
 | 
117  | 
    "{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
 | 
118  | 
    "{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 | 119  | 
by (auto elim: F.not_neutral_contains_not_neutral)  | 
| 
66804
 
3f9bb52082c4
avoid name clashes on interpretation of abstract locales
 
haftmann 
parents: 
64272 
diff
changeset
 | 
120  | 
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
 | 
121  | 
    "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
 | 
122  | 
      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
 | 
123  | 
  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
 | 
124  | 
    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
 | 
125  | 
    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
 | 
126  | 
      expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"])
 | 
| 58197 | 127  | 
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
 | 
128  | 
    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
 | 
129  | 
      expand_superset [of "{a. \<exists>b. g a b \<noteq> \<^bold>1}"])
 | 
| 58197 | 130  | 
qed  | 
131  | 
||
132  | 
lemma cartesian_product:  | 
|
133  | 
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
 | 
134  | 
  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 | 135  | 
shows "G (\<lambda>a. G (g a)) = G (\<lambda>(a, b). g a b)"  | 
136  | 
proof -  | 
|
| 60500 | 137  | 
from subset \<open>finite C\<close> have fin_prod: "finite (?A \<times> ?B)"  | 
| 58197 | 138  | 
by (rule finite_subset)  | 
139  | 
from fin_prod have "finite ?A" and "finite ?B"  | 
|
140  | 
by (auto simp add: finite_cartesian_product_iff)  | 
|
141  | 
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
 | 
142  | 
    (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})"
 | 
| 58197 | 143  | 
apply (subst expand_superset [of "?B"])  | 
| 60500 | 144  | 
apply (rule \<open>finite ?B\<close>)  | 
| 58197 | 145  | 
apply auto  | 
146  | 
apply (subst expand_superset [of "?A"])  | 
|
| 60500 | 147  | 
apply (rule \<open>finite ?A\<close>)  | 
| 58197 | 148  | 
apply auto  | 
149  | 
apply (erule F.not_neutral_contains_not_neutral)  | 
|
150  | 
apply auto  | 
|
151  | 
done  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
61955 
diff
changeset
 | 
152  | 
  have "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> ?A \<times> ?B"
 | 
| 58197 | 153  | 
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
 | 
154  | 
  with subset have **: "{p. (case p of (a, b) \<Rightarrow> g a b) \<noteq> \<^bold>1} \<subseteq> C"
 | 
| 58197 | 155  | 
by blast  | 
156  | 
show ?thesis  | 
|
157  | 
apply (simp add: *)  | 
|
158  | 
apply (simp add: F.cartesian_product)  | 
|
159  | 
apply (subst expand_superset [of C])  | 
|
| 60500 | 160  | 
apply (rule \<open>finite C\<close>)  | 
| 58197 | 161  | 
apply (simp_all add: **)  | 
162  | 
apply (rule F.same_carrierI [of C])  | 
|
| 60500 | 163  | 
apply (rule \<open>finite C\<close>)  | 
| 58197 | 164  | 
apply (simp_all add: subset)  | 
165  | 
apply auto  | 
|
166  | 
done  | 
|
167  | 
qed  | 
|
168  | 
||
169  | 
lemma cartesian_product2:  | 
|
170  | 
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
 | 
171  | 
  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 | 172  | 
shows "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>(a, b, c). g a b c)"  | 
173  | 
proof -  | 
|
174  | 
have bij: "bij (\<lambda>(a, b, c). ((a, b), c))"  | 
|
175  | 
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
 | 
176  | 
  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
 | 
177  | 
by auto (insert subset, blast)  | 
| 58197 | 178  | 
with fin have "G (\<lambda>p. G (g (fst p) (snd p))) = G (\<lambda>(p, c). g (fst p) (snd p) c)"  | 
179  | 
by (rule cartesian_product)  | 
|
180  | 
then have "G (\<lambda>(a, b). G (g a b)) = G (\<lambda>((a, b), c). g a b c)"  | 
|
181  | 
by (auto simp add: split_def)  | 
|
182  | 
also have "G (\<lambda>((a, b), c). g a b c) = G (\<lambda>(a, b, c). g a b c)"  | 
|
183  | 
using bij by (rule reindex_cong [of "\<lambda>(a, b, c). ((a, b), c)"]) (simp add: fun_eq_iff)  | 
|
184  | 
finally show ?thesis .  | 
|
185  | 
qed  | 
|
186  | 
||
187  | 
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
 | 
188  | 
"G (\<lambda>b. if b = a then g b else \<^bold>1) = g a"  | 
| 58197 | 189  | 
proof -  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
61955 
diff
changeset
 | 
190  | 
  have "{b. (if b = a then g b else \<^bold>1) \<noteq> \<^bold>1} \<subseteq> {a}" by auto
 | 
| 58197 | 191  | 
  then show ?thesis by (simp add: expand_superset [of "{a}"])
 | 
192  | 
qed  | 
|
193  | 
||
194  | 
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
 | 
195  | 
"G (\<lambda>b. if a = b then g b else \<^bold>1) = g a"  | 
| 58197 | 196  | 
proof -  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
61955 
diff
changeset
 | 
197  | 
have "(\<lambda>b. if a = b then g b else \<^bold>1) = (\<lambda>b. if b = a then g b else \<^bold>1)"  | 
| 58197 | 198  | 
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
 | 
199  | 
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 | 200  | 
by (simp cong del: cong)  | 
| 58197 | 201  | 
then show ?thesis by simp  | 
202  | 
qed  | 
|
203  | 
||
204  | 
end  | 
|
205  | 
||
206  | 
||
207  | 
subsection \<open>Concrete sum\<close>  | 
|
208  | 
||
209  | 
context comm_monoid_add  | 
|
210  | 
begin  | 
|
211  | 
||
| 61776 | 212  | 
sublocale Sum_any: comm_monoid_fun plus 0  | 
| 67764 | 213  | 
rewrites "comm_monoid_set.F plus 0 = sum"  | 
| 63433 | 214  | 
defines Sum_any = Sum_any.G  | 
| 58197 | 215  | 
proof -  | 
216  | 
show "comm_monoid_fun plus 0" ..  | 
|
| 61605 | 217  | 
then interpret Sum_any: comm_monoid_fun plus 0 .  | 
| 64267 | 218  | 
from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym)  | 
| 58197 | 219  | 
qed  | 
220  | 
||
221  | 
end  | 
|
222  | 
||
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61776 
diff
changeset
 | 
223  | 
syntax (ASCII)  | 
| 58197 | 224  | 
  "_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
 | 
225  | 
syntax  | 
| 58197 | 226  | 
  "_Sum_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"    ("(3\<Sum>_. _)" [0, 10] 10)
 | 
227  | 
translations  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61776 
diff
changeset
 | 
228  | 
"\<Sum>a. b" \<rightleftharpoons> "CONST Sum_any (\<lambda>a. b)"  | 
| 58197 | 229  | 
|
230  | 
lemma Sum_any_left_distrib:  | 
|
231  | 
fixes r :: "'a :: semiring_0"  | 
|
232  | 
  assumes "finite {a. g a \<noteq> 0}"
 | 
|
233  | 
shows "Sum_any g * r = (\<Sum>n. g n * r)"  | 
|
234  | 
proof -  | 
|
235  | 
note assms  | 
|
236  | 
  moreover have "{a. g a * r \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
 | 
|
237  | 
ultimately show ?thesis  | 
|
| 64267 | 238  | 
    by (simp add: sum_distrib_right Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
 | 
| 58197 | 239  | 
qed  | 
240  | 
||
241  | 
lemma Sum_any_right_distrib:  | 
|
242  | 
fixes r :: "'a :: semiring_0"  | 
|
243  | 
  assumes "finite {a. g a \<noteq> 0}"
 | 
|
244  | 
shows "r * Sum_any g = (\<Sum>n. r * g n)"  | 
|
245  | 
proof -  | 
|
246  | 
note assms  | 
|
247  | 
  moreover have "{a. r * g a \<noteq> 0} \<subseteq> {a. g a \<noteq> 0}" by auto
 | 
|
248  | 
ultimately show ?thesis  | 
|
| 64267 | 249  | 
    by (simp add: sum_distrib_left Sum_any.expand_superset [of "{a. g a \<noteq> 0}"])
 | 
| 58197 | 250  | 
qed  | 
251  | 
||
252  | 
lemma Sum_any_product:  | 
|
253  | 
fixes f g :: "'b \<Rightarrow> 'a::semiring_0"  | 
|
254  | 
  assumes "finite {a. f a \<noteq> 0}" and "finite {b. g b \<noteq> 0}"
 | 
|
255  | 
shows "Sum_any f * Sum_any g = (\<Sum>a. \<Sum>b. f a * g b)"  | 
|
256  | 
proof -  | 
|
257  | 
  have subset_f: "{a. (\<Sum>b. f a * g b) \<noteq> 0} \<subseteq> {a. f a \<noteq> 0}"
 | 
|
258  | 
by rule (simp, rule, auto)  | 
|
259  | 
  moreover have subset_g: "\<And>a. {b. f a * g b \<noteq> 0} \<subseteq> {b. g b \<noteq> 0}"
 | 
|
260  | 
by rule (simp, rule, auto)  | 
|
261  | 
ultimately show ?thesis using assms  | 
|
262  | 
by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g]  | 
|
263  | 
      Sum_any.expand_superset [of "{a. f a \<noteq> 0}"] Sum_any.expand_superset [of "{b. g b \<noteq> 0}"]
 | 
|
| 64267 | 264  | 
sum_product)  | 
| 58197 | 265  | 
qed  | 
266  | 
||
| 58437 | 267  | 
lemma Sum_any_eq_zero_iff [simp]:  | 
268  | 
fixes f :: "'a \<Rightarrow> nat"  | 
|
269  | 
  assumes "finite {a. f a \<noteq> 0}"
 | 
|
270  | 
shows "Sum_any f = 0 \<longleftrightarrow> f = (\<lambda>_. 0)"  | 
|
271  | 
using assms by (simp add: Sum_any.expand_set fun_eq_iff)  | 
|
272  | 
||
| 58197 | 273  | 
|
274  | 
subsection \<open>Concrete product\<close>  | 
|
275  | 
||
276  | 
context comm_monoid_mult  | 
|
277  | 
begin  | 
|
278  | 
||
| 61776 | 279  | 
sublocale Prod_any: comm_monoid_fun times 1  | 
| 67764 | 280  | 
rewrites "comm_monoid_set.F times 1 = prod"  | 
| 63433 | 281  | 
defines Prod_any = Prod_any.G  | 
| 58197 | 282  | 
proof -  | 
283  | 
show "comm_monoid_fun times 1" ..  | 
|
| 61605 | 284  | 
then interpret Prod_any: comm_monoid_fun times 1 .  | 
| 64272 | 285  | 
from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym)  | 
| 58197 | 286  | 
qed  | 
287  | 
||
288  | 
end  | 
|
289  | 
||
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61776 
diff
changeset
 | 
290  | 
syntax (ASCII)  | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61776 
diff
changeset
 | 
291  | 
  "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _. _)" [0, 10] 10)
 | 
| 58197 | 292  | 
syntax  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61776 
diff
changeset
 | 
293  | 
  "_Prod_any" :: "pttrn \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_. _)" [0, 10] 10)
 | 
| 58197 | 294  | 
translations  | 
295  | 
"\<Prod>a. b" == "CONST Prod_any (\<lambda>a. b)"  | 
|
296  | 
||
297  | 
lemma Prod_any_zero:  | 
|
298  | 
fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1"  | 
|
299  | 
  assumes "finite {a. f a \<noteq> 1}"
 | 
|
300  | 
assumes "f a = 0"  | 
|
301  | 
shows "(\<Prod>a. f a) = 0"  | 
|
302  | 
proof -  | 
|
| 60500 | 303  | 
from \<open>f a = 0\<close> have "f a \<noteq> 1" by simp  | 
304  | 
with \<open>f a = 0\<close> have "\<exists>a. f a \<noteq> 1 \<and> f a = 0" by blast  | 
|
305  | 
  with \<open>finite {a. f a \<noteq> 1}\<close> show ?thesis
 | 
|
| 64272 | 306  | 
by (simp add: Prod_any.expand_set prod_zero)  | 
| 58197 | 307  | 
qed  | 
308  | 
||
309  | 
lemma Prod_any_not_zero:  | 
|
310  | 
fixes f :: "'b \<Rightarrow> 'a :: comm_semiring_1"  | 
|
311  | 
  assumes "finite {a. f a \<noteq> 1}"
 | 
|
312  | 
assumes "(\<Prod>a. f a) \<noteq> 0"  | 
|
313  | 
shows "f a \<noteq> 0"  | 
|
314  | 
using assms Prod_any_zero [of f] by blast  | 
|
315  | 
||
| 58437 | 316  | 
lemma power_Sum_any:  | 
317  | 
  assumes "finite {a. f a \<noteq> 0}"
 | 
|
318  | 
shows "c ^ (\<Sum>a. f a) = (\<Prod>a. c ^ f a)"  | 
|
319  | 
proof -  | 
|
320  | 
  have "{a. c ^ f a \<noteq> 1} \<subseteq> {a. f a \<noteq> 0}"
 | 
|
321  | 
by (auto intro: ccontr)  | 
|
322  | 
with assms show ?thesis  | 
|
| 64267 | 323  | 
by (simp add: Sum_any.expand_set Prod_any.expand_superset power_sum)  | 
| 58437 | 324  | 
qed  | 
325  | 
||
| 58197 | 326  | 
end  |