author | wenzelm |
Mon, 03 Feb 2025 20:22:51 +0100 (3 months ago) | |
changeset 82073 | 879be333e939 |
parent 81438 | 95c9af7483b1 |
permissions | -rw-r--r-- |
14706 | 1 |
(* Title: HOL/Algebra/FiniteProduct.thy |
13936 | 2 |
Author: Clemens Ballarin, started 19 November 2002 |
3 |
||
4 |
This file is largely based on HOL/Finite_Set.thy. |
|
5 |
*) |
|
6 |
||
35849 | 7 |
theory FiniteProduct |
8 |
imports Group |
|
9 |
begin |
|
13936 | 10 |
|
61382 | 11 |
subsection \<open>Product Operator for Commutative Monoids\<close> |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
12 |
|
61382 | 13 |
subsubsection \<open>Inductive Definition of a Relation for Products over Sets\<close> |
13936 | 14 |
|
63167 | 15 |
text \<open>Instantiation of locale \<open>LC\<close> of theory \<open>Finite_Set\<close> is not |
68458 | 16 |
possible, because here we have explicit typing rules like |
63167 | 17 |
\<open>x \<in> carrier G\<close>. We introduce an explicit argument for the domain |
18 |
\<open>D\<close>.\<close> |
|
13936 | 19 |
|
23746 | 20 |
inductive_set |
68458 | 21 |
foldSetD :: "['a set, 'b \<Rightarrow> 'a \<Rightarrow> 'a, 'a] \<Rightarrow> ('b set * 'a) set" |
22 |
for D :: "'a set" and f :: "'b \<Rightarrow> 'a \<Rightarrow> 'a" and e :: 'a |
|
23746 | 23 |
where |
68458 | 24 |
emptyI [intro]: "e \<in> D \<Longrightarrow> ({}, e) \<in> foldSetD D f e" |
25 |
| insertI [intro]: "\<lbrakk>x \<notin> A; f x y \<in> D; (A, y) \<in> foldSetD D f e\<rbrakk> \<Longrightarrow> |
|
14750 | 26 |
(insert x A, f x y) \<in> foldSetD D f e" |
13936 | 27 |
|
14750 | 28 |
inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> foldSetD D f e" |
13936 | 29 |
|
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
30 |
definition |
68458 | 31 |
foldD :: "['a set, 'b \<Rightarrow> 'a \<Rightarrow> 'a, 'a, 'b set] \<Rightarrow> 'a" |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
32 |
where "foldD D f e A = (THE x. (A, x) \<in> foldSetD D f e)" |
13936 | 33 |
|
68517 | 34 |
lemma foldSetD_closed: "(A, z) \<in> foldSetD D f e \<Longrightarrow> z \<in> D" |
23746 | 35 |
by (erule foldSetD.cases) auto |
13936 | 36 |
|
37 |
lemma Diff1_foldSetD: |
|
68458 | 38 |
"\<lbrakk>(A - {x}, y) \<in> foldSetD D f e; x \<in> A; f x y \<in> D\<rbrakk> \<Longrightarrow> |
14750 | 39 |
(A, f x y) \<in> foldSetD D f e" |
68458 | 40 |
by (metis Diff_insert_absorb foldSetD.insertI mk_disjoint_insert) |
13936 | 41 |
|
68458 | 42 |
lemma foldSetD_imp_finite [simp]: "(A, x) \<in> foldSetD D f e \<Longrightarrow> finite A" |
13936 | 43 |
by (induct set: foldSetD) auto |
44 |
||
45 |
lemma finite_imp_foldSetD: |
|
68458 | 46 |
"\<lbrakk>finite A; e \<in> D; \<And>x y. \<lbrakk>x \<in> A; y \<in> D\<rbrakk> \<Longrightarrow> f x y \<in> D\<rbrakk> |
47 |
\<Longrightarrow> \<exists>x. (A, x) \<in> foldSetD D f e" |
|
22265 | 48 |
proof (induct set: finite) |
13936 | 49 |
case empty then show ?case by auto |
50 |
next |
|
15328 | 51 |
case (insert x F) |
14750 | 52 |
then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto |
53 |
with insert have "y \<in> D" by (auto dest: foldSetD_closed) |
|
54 |
with y and insert have "(insert x F, f x y) \<in> foldSetD D f e" |
|
13936 | 55 |
by (intro foldSetD.intros) auto |
56 |
then show ?case .. |
|
57 |
qed |
|
58 |
||
68517 | 59 |
lemma foldSetD_backwards: |
60 |
assumes "A \<noteq> {}" "(A, z) \<in> foldSetD D f e" |
|
61 |
shows "\<exists>x y. x \<in> A \<and> (A - { x }, y) \<in> foldSetD D f e \<and> z = f x y" |
|
62 |
using assms(2) by (cases) (simp add: assms(1), metis Diff_insert_absorb insertI1) |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
63 |
|
68517 | 64 |
subsubsection \<open>Left-Commutative Operations\<close> |
13936 | 65 |
|
66 |
locale LCD = |
|
67 |
fixes B :: "'b set" |
|
68 |
and D :: "'a set" |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
69 |
and f :: "'b \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<cdot>\<close> 70) |
13936 | 70 |
assumes left_commute: |
68458 | 71 |
"\<lbrakk>x \<in> B; y \<in> B; z \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
72 |
and f_closed [simp, intro!]: "!!x y. \<lbrakk>x \<in> B; y \<in> D\<rbrakk> \<Longrightarrow> f x y \<in> D" |
|
13936 | 73 |
|
68458 | 74 |
lemma (in LCD) foldSetD_closed [dest]: "(A, z) \<in> foldSetD D f e \<Longrightarrow> z \<in> D" |
23746 | 75 |
by (erule foldSetD.cases) auto |
13936 | 76 |
|
77 |
lemma (in LCD) Diff1_foldSetD: |
|
68458 | 78 |
"\<lbrakk>(A - {x}, y) \<in> foldSetD D f e; x \<in> A; A \<subseteq> B\<rbrakk> \<Longrightarrow> |
14750 | 79 |
(A, f x y) \<in> foldSetD D f e" |
68458 | 80 |
by (meson Diff1_foldSetD f_closed local.foldSetD_closed subsetCE) |
13936 | 81 |
|
82 |
lemma (in LCD) finite_imp_foldSetD: |
|
68458 | 83 |
"\<lbrakk>finite A; A \<subseteq> B; e \<in> D\<rbrakk> \<Longrightarrow> \<exists>x. (A, x) \<in> foldSetD D f e" |
22265 | 84 |
proof (induct set: finite) |
13936 | 85 |
case empty then show ?case by auto |
86 |
next |
|
15328 | 87 |
case (insert x F) |
14750 | 88 |
then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto |
89 |
with insert have "y \<in> D" by auto |
|
90 |
with y and insert have "(insert x F, f x y) \<in> foldSetD D f e" |
|
13936 | 91 |
by (intro foldSetD.intros) auto |
92 |
then show ?case .. |
|
93 |
qed |
|
94 |
||
68458 | 95 |
|
13936 | 96 |
lemma (in LCD) foldSetD_determ_aux: |
68458 | 97 |
assumes "e \<in> D" and A: "card A < n" "A \<subseteq> B" "(A, x) \<in> foldSetD D f e" "(A, y) \<in> foldSetD D f e" |
98 |
shows "y = x" |
|
99 |
using A |
|
100 |
proof (induction n arbitrary: A x y) |
|
101 |
case 0 |
|
102 |
then show ?case |
|
103 |
by auto |
|
104 |
next |
|
105 |
case (Suc n) |
|
106 |
then consider "card A = n" | "card A < n" |
|
107 |
by linarith |
|
108 |
then show ?case |
|
109 |
proof cases |
|
110 |
case 1 |
|
111 |
show ?thesis |
|
112 |
using foldSetD.cases [OF \<open>(A,x) \<in> foldSetD D (\<cdot>) e\<close>] |
|
113 |
proof cases |
|
114 |
case 1 |
|
115 |
then show ?thesis |
|
116 |
using \<open>(A,y) \<in> foldSetD D (\<cdot>) e\<close> by auto |
|
117 |
next |
|
81438 | 118 |
case A': (2 x' A' y') |
68458 | 119 |
show ?thesis |
120 |
using foldSetD.cases [OF \<open>(A,y) \<in> foldSetD D (\<cdot>) e\<close>] |
|
121 |
proof cases |
|
122 |
case 1 |
|
123 |
then show ?thesis |
|
124 |
using \<open>(A,x) \<in> foldSetD D (\<cdot>) e\<close> by auto |
|
125 |
next |
|
81438 | 126 |
case A'': (2 x'' A'' y'') |
68458 | 127 |
show ?thesis |
128 |
proof (cases "x' = x''") |
|
129 |
case True |
|
130 |
show ?thesis |
|
131 |
proof (cases "y' = y''") |
|
132 |
case True |
|
133 |
then show ?thesis |
|
134 |
using A' A'' \<open>x' = x''\<close> by (blast elim!: equalityE) |
|
135 |
next |
|
136 |
case False |
|
137 |
then show ?thesis |
|
138 |
using A' A'' \<open>x' = x''\<close> |
|
139 |
by (metis \<open>card A = n\<close> Suc.IH Suc.prems(2) card_insert_disjoint foldSetD_imp_finite insert_eq_iff insert_subset lessI) |
|
140 |
qed |
|
141 |
next |
|
142 |
case False |
|
143 |
then have *: "A' - {x''} = A'' - {x'}" "x'' \<in> A'" "x' \<in> A''" |
|
144 |
using A' A'' by fastforce+ |
|
145 |
then have "A' = insert x'' A'' - {x'}" |
|
146 |
using \<open>x' \<notin> A'\<close> by blast |
|
147 |
then have card: "card A' \<le> card A''" |
|
148 |
using A' A'' * by (metis card_Suc_Diff1 eq_refl foldSetD_imp_finite) |
|
149 |
obtain u where u: "(A' - {x''}, u) \<in> foldSetD D (\<cdot>) e" |
|
150 |
using finite_imp_foldSetD [of "A' - {x''}"] A' Diff_insert \<open>A \<subseteq> B\<close> \<open>e \<in> D\<close> by fastforce |
|
151 |
have "y' = f x'' u" |
|
152 |
using Diff1_foldSetD [OF u] \<open>x'' \<in> A'\<close> \<open>card A = n\<close> A' Suc.IH \<open>A \<subseteq> B\<close> by auto |
|
153 |
then have "(A'' - {x'}, u) \<in> foldSetD D f e" |
|
154 |
using "*"(1) u by auto |
|
155 |
then have "y'' = f x' u" |
|
156 |
using A'' by (metis * \<open>card A = n\<close> A'(1) Diff1_foldSetD Suc.IH \<open>A \<subseteq> B\<close> |
|
157 |
card card_Suc_Diff1 card_insert_disjoint foldSetD_imp_finite insert_subset le_imp_less_Suc) |
|
158 |
then show ?thesis |
|
159 |
using A' A'' |
|
160 |
by (metis \<open>A \<subseteq> B\<close> \<open>y' = x'' \<cdot> u\<close> insert_subset left_commute local.foldSetD_closed u) |
|
161 |
qed |
|
162 |
qed |
|
163 |
qed |
|
164 |
next |
|
165 |
case 2 with Suc show ?thesis by blast |
|
166 |
qed |
|
167 |
qed |
|
13936 | 168 |
|
169 |
lemma (in LCD) foldSetD_determ: |
|
68458 | 170 |
"\<lbrakk>(A, x) \<in> foldSetD D f e; (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B\<rbrakk> |
171 |
\<Longrightarrow> y = x" |
|
13936 | 172 |
by (blast intro: foldSetD_determ_aux [rule_format]) |
173 |
||
174 |
lemma (in LCD) foldD_equality: |
|
68458 | 175 |
"\<lbrakk>(A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B\<rbrakk> \<Longrightarrow> foldD D f e A = y" |
13936 | 176 |
by (unfold foldD_def) (blast intro: foldSetD_determ) |
177 |
||
178 |
lemma foldD_empty [simp]: |
|
68458 | 179 |
"e \<in> D \<Longrightarrow> foldD D f e {} = e" |
13936 | 180 |
by (unfold foldD_def) blast |
181 |
||
182 |
lemma (in LCD) foldD_insert_aux: |
|
68458 | 183 |
"\<lbrakk>x \<notin> A; x \<in> B; e \<in> D; A \<subseteq> B\<rbrakk> |
184 |
\<Longrightarrow> ((insert x A, v) \<in> foldSetD D f e) \<longleftrightarrow> (\<exists>y. (A, y) \<in> foldSetD D f e \<and> v = f x y)" |
|
13936 | 185 |
apply auto |
68458 | 186 |
by (metis Diff_insert_absorb f_closed finite_Diff foldSetD.insertI foldSetD_determ foldSetD_imp_finite insert_subset local.finite_imp_foldSetD local.foldSetD_closed) |
13936 | 187 |
|
188 |
lemma (in LCD) foldD_insert: |
|
68458 | 189 |
assumes "finite A" "x \<notin> A" "x \<in> B" "e \<in> D" "A \<subseteq> B" |
190 |
shows "foldD D f e (insert x A) = f x (foldD D f e A)" |
|
191 |
proof - |
|
192 |
have "(THE v. \<exists>y. (A, y) \<in> foldSetD D (\<cdot>) e \<and> v = x \<cdot> y) = x \<cdot> (THE y. (A, y) \<in> foldSetD D (\<cdot>) e)" |
|
193 |
by (rule the_equality) (use assms foldD_def foldD_equality foldD_def finite_imp_foldSetD in \<open>metis+\<close>) |
|
194 |
then show ?thesis |
|
195 |
unfolding foldD_def using assms by (simp add: foldD_insert_aux) |
|
196 |
qed |
|
13936 | 197 |
|
198 |
lemma (in LCD) foldD_closed [simp]: |
|
68458 | 199 |
"\<lbrakk>finite A; e \<in> D; A \<subseteq> B\<rbrakk> \<Longrightarrow> foldD D f e A \<in> D" |
22265 | 200 |
proof (induct set: finite) |
46721 | 201 |
case empty then show ?case by simp |
13936 | 202 |
next |
203 |
case insert then show ?case by (simp add: foldD_insert) |
|
204 |
qed |
|
205 |
||
206 |
lemma (in LCD) foldD_commute: |
|
68458 | 207 |
"\<lbrakk>finite A; x \<in> B; e \<in> D; A \<subseteq> B\<rbrakk> \<Longrightarrow> |
13936 | 208 |
f x (foldD D f e A) = foldD D f (f x e) A" |
68458 | 209 |
by (induct set: finite) (auto simp add: left_commute foldD_insert) |
13936 | 210 |
|
211 |
lemma Int_mono2: |
|
68458 | 212 |
"\<lbrakk>A \<subseteq> C; B \<subseteq> C\<rbrakk> \<Longrightarrow> A Int B \<subseteq> C" |
13936 | 213 |
by blast |
214 |
||
215 |
lemma (in LCD) foldD_nest_Un_Int: |
|
68458 | 216 |
"\<lbrakk>finite A; finite C; e \<in> D; A \<subseteq> B; C \<subseteq> B\<rbrakk> \<Longrightarrow> |
13936 | 217 |
foldD D f (foldD D f e C) A = foldD D f (foldD D f e (A Int C)) (A Un C)" |
68458 | 218 |
proof (induction set: finite) |
219 |
case (insert x F) |
|
220 |
then show ?case |
|
221 |
by (simp add: foldD_insert foldD_commute Int_insert_left insert_absorb Int_mono2) |
|
222 |
qed simp |
|
13936 | 223 |
|
224 |
lemma (in LCD) foldD_nest_Un_disjoint: |
|
68458 | 225 |
"\<lbrakk>finite A; finite B; A Int B = {}; e \<in> D; A \<subseteq> B; C \<subseteq> B\<rbrakk> |
226 |
\<Longrightarrow> foldD D f e (A Un B) = foldD D f (foldD D f e B) A" |
|
13936 | 227 |
by (simp add: foldD_nest_Un_Int) |
228 |
||
63167 | 229 |
\<comment> \<open>Delete rules to do with \<open>foldSetD\<close> relation.\<close> |
13936 | 230 |
|
231 |
declare foldSetD_imp_finite [simp del] |
|
232 |
empty_foldSetDE [rule del] |
|
233 |
foldSetD.intros [rule del] |
|
234 |
declare (in LCD) |
|
235 |
foldSetD_closed [rule del] |
|
236 |
||
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
237 |
|
61382 | 238 |
text \<open>Commutative Monoids\<close> |
13936 | 239 |
|
61382 | 240 |
text \<open> |
68458 | 241 |
We enter a more restrictive context, with \<open>f :: 'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> |
242 |
instead of \<open>'b \<Rightarrow> 'a \<Rightarrow> 'a\<close>. |
|
61382 | 243 |
\<close> |
13936 | 244 |
|
245 |
locale ACeD = |
|
246 |
fixes D :: "'a set" |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
247 |
and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<cdot>\<close> 70) |
13936 | 248 |
and e :: 'a |
68458 | 249 |
assumes ident [simp]: "x \<in> D \<Longrightarrow> x \<cdot> e = x" |
250 |
and commute: "\<lbrakk>x \<in> D; y \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> y = y \<cdot> x" |
|
251 |
and assoc: "\<lbrakk>x \<in> D; y \<in> D; z \<in> D\<rbrakk> \<Longrightarrow> (x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
14750 | 252 |
and e_closed [simp]: "e \<in> D" |
68458 | 253 |
and f_closed [simp]: "\<lbrakk>x \<in> D; y \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> y \<in> D" |
13936 | 254 |
|
255 |
lemma (in ACeD) left_commute: |
|
68458 | 256 |
"\<lbrakk>x \<in> D; y \<in> D; z \<in> D\<rbrakk> \<Longrightarrow> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
13936 | 257 |
proof - |
14750 | 258 |
assume D: "x \<in> D" "y \<in> D" "z \<in> D" |
13936 | 259 |
then have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp add: commute) |
260 |
also from D have "... = y \<cdot> (z \<cdot> x)" by (simp add: assoc) |
|
261 |
also from D have "z \<cdot> x = x \<cdot> z" by (simp add: commute) |
|
262 |
finally show ?thesis . |
|
263 |
qed |
|
264 |
||
265 |
lemmas (in ACeD) AC = assoc commute left_commute |
|
266 |
||
68458 | 267 |
lemma (in ACeD) left_ident [simp]: "x \<in> D \<Longrightarrow> e \<cdot> x = x" |
13936 | 268 |
proof - |
23350 | 269 |
assume "x \<in> D" |
270 |
then have "x \<cdot> e = x" by (rule ident) |
|
61382 | 271 |
with \<open>x \<in> D\<close> show ?thesis by (simp add: commute) |
13936 | 272 |
qed |
273 |
||
274 |
lemma (in ACeD) foldD_Un_Int: |
|
68458 | 275 |
"\<lbrakk>finite A; finite B; A \<subseteq> D; B \<subseteq> D\<rbrakk> \<Longrightarrow> |
13936 | 276 |
foldD D f e A \<cdot> foldD D f e B = |
277 |
foldD D f e (A Un B) \<cdot> foldD D f e (A Int B)" |
|
68458 | 278 |
proof (induction set: finite) |
279 |
case empty |
|
280 |
then show ?case |
|
281 |
by(simp add: left_commute LCD.foldD_closed [OF LCD.intro [of D]]) |
|
282 |
next |
|
283 |
case (insert x F) |
|
284 |
then show ?case |
|
285 |
by(simp add: AC insert_absorb Int_insert_left Int_mono2 |
|
286 |
LCD.foldD_insert [OF LCD.intro [of D]] |
|
287 |
LCD.foldD_closed [OF LCD.intro [of D]]) |
|
288 |
qed |
|
13936 | 289 |
|
290 |
lemma (in ACeD) foldD_Un_disjoint: |
|
68458 | 291 |
"\<lbrakk>finite A; finite B; A Int B = {}; A \<subseteq> D; B \<subseteq> D\<rbrakk> \<Longrightarrow> |
13936 | 292 |
foldD D f e (A Un B) = foldD D f e A \<cdot> foldD D f e B" |
293 |
by (simp add: foldD_Un_Int |
|
32693 | 294 |
left_commute LCD.foldD_closed [OF LCD.intro [of D]]) |
13936 | 295 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
296 |
|
61382 | 297 |
subsubsection \<open>Products over Finite Sets\<close> |
13936 | 298 |
|
35847 | 299 |
definition |
68458 | 300 |
finprod :: "[('b, 'm) monoid_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow> 'b" |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
301 |
where "finprod G f A = |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
302 |
(if finite A |
67091 | 303 |
then foldD (carrier G) (mult G \<circ> f) \<one>\<^bsub>G\<^esub> A |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
304 |
else \<one>\<^bsub>G\<^esub>)" |
13936 | 305 |
|
14651 | 306 |
syntax |
68458 | 307 |
"_finprod" :: "index \<Rightarrow> idt \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'b" |
81142 | 308 |
(\<open>(\<open>indent=3 notation=\<open>binder \<Otimes>\<close>\<close>\<Otimes>__\<in>_. _)\<close> [1000, 0, 51, 10] 10) |
80768 | 309 |
syntax_consts |
310 |
"_finprod" \<rightleftharpoons> finprod |
|
14651 | 311 |
translations |
62105 | 312 |
"\<Otimes>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finprod G (%i. b) A" |
63167 | 313 |
\<comment> \<open>Beware of argument permutation!\<close> |
13936 | 314 |
|
68458 | 315 |
lemma (in comm_monoid) finprod_empty [simp]: |
13936 | 316 |
"finprod G f {} = \<one>" |
317 |
by (simp add: finprod_def) |
|
318 |
||
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
319 |
lemma (in comm_monoid) finprod_infinite[simp]: |
68458 | 320 |
"\<not> finite A \<Longrightarrow> finprod G f A = \<one>" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
321 |
by (simp add: finprod_def) |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
322 |
|
13936 | 323 |
declare funcsetI [intro] |
324 |
funcset_mem [dest] |
|
325 |
||
27933 | 326 |
context comm_monoid begin |
327 |
||
328 |
lemma finprod_insert [simp]: |
|
68458 | 329 |
assumes "finite F" "a \<notin> F" "f \<in> F \<rightarrow> carrier G" "f a \<in> carrier G" |
330 |
shows "finprod G f (insert a F) = f a \<otimes> finprod G f F" |
|
331 |
proof - |
|
332 |
have "finprod G f (insert a F) = foldD (carrier G) ((\<otimes>) \<circ> f) \<one> (insert a F)" |
|
333 |
by (simp add: finprod_def assms) |
|
334 |
also have "... = ((\<otimes>) \<circ> f) a (foldD (carrier G) ((\<otimes>) \<circ> f) \<one> F)" |
|
335 |
by (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]]) |
|
336 |
(use assms in \<open>auto simp: m_lcomm Pi_iff\<close>) |
|
337 |
also have "... = f a \<otimes> finprod G f F" |
|
338 |
using \<open>finite F\<close> by (auto simp add: finprod_def) |
|
339 |
finally show ?thesis . |
|
340 |
qed |
|
13936 | 341 |
|
68447
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
342 |
lemma finprod_one_eqI: "(\<And>x. x \<in> A \<Longrightarrow> f x = \<one>) \<Longrightarrow> finprod G f A = \<one>" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
343 |
proof (induct A rule: infinite_finite_induct) |
13936 | 344 |
case empty show ?case by simp |
345 |
next |
|
15328 | 346 |
case (insert a A) |
68447
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
347 |
have "(\<lambda>i. \<one>) \<in> A \<rightarrow> carrier G" by auto |
13936 | 348 |
with insert show ?case by simp |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
349 |
qed simp |
13936 | 350 |
|
68447
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
351 |
lemma finprod_one [simp]: "(\<Otimes>i\<in>A. \<one>) = \<one>" |
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
352 |
by (simp add: finprod_one_eqI) |
0beb927eed89
Adjusting Number_Theory for new Algebra
paulson <lp15@cam.ac.uk>
parents:
68445
diff
changeset
|
353 |
|
27933 | 354 |
lemma finprod_closed [simp]: |
13936 | 355 |
fixes A |
68458 | 356 |
assumes f: "f \<in> A \<rightarrow> carrier G" |
13936 | 357 |
shows "finprod G f A \<in> carrier G" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
358 |
using f |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
359 |
proof (induct A rule: infinite_finite_induct) |
13936 | 360 |
case empty show ?case by simp |
361 |
next |
|
15328 | 362 |
case (insert a A) |
13936 | 363 |
then have a: "f a \<in> carrier G" by fast |
61384 | 364 |
from insert have A: "f \<in> A \<rightarrow> carrier G" by fast |
13936 | 365 |
from insert A a show ?case by simp |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
366 |
qed simp |
13936 | 367 |
|
368 |
lemma funcset_Int_left [simp, intro]: |
|
68458 | 369 |
"\<lbrakk>f \<in> A \<rightarrow> C; f \<in> B \<rightarrow> C\<rbrakk> \<Longrightarrow> f \<in> A Int B \<rightarrow> C" |
13936 | 370 |
by fast |
371 |
||
372 |
lemma funcset_Un_left [iff]: |
|
67091 | 373 |
"(f \<in> A Un B \<rightarrow> C) = (f \<in> A \<rightarrow> C \<and> f \<in> B \<rightarrow> C)" |
13936 | 374 |
by fast |
375 |
||
27933 | 376 |
lemma finprod_Un_Int: |
68458 | 377 |
"\<lbrakk>finite A; finite B; g \<in> A \<rightarrow> carrier G; g \<in> B \<rightarrow> carrier G\<rbrakk> \<Longrightarrow> |
13936 | 378 |
finprod G g (A Un B) \<otimes> finprod G g (A Int B) = |
379 |
finprod G g A \<otimes> finprod G g B" |
|
63167 | 380 |
\<comment> \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close> |
22265 | 381 |
proof (induct set: finite) |
46721 | 382 |
case empty then show ?case by simp |
13936 | 383 |
next |
15328 | 384 |
case (insert a A) |
13936 | 385 |
then have a: "g a \<in> carrier G" by fast |
61384 | 386 |
from insert have A: "g \<in> A \<rightarrow> carrier G" by fast |
13936 | 387 |
from insert A a show ?case |
68458 | 388 |
by (simp add: m_ac Int_insert_left insert_absorb Int_mono2) |
13936 | 389 |
qed |
390 |
||
27933 | 391 |
lemma finprod_Un_disjoint: |
68458 | 392 |
"\<lbrakk>finite A; finite B; A Int B = {}; |
393 |
g \<in> A \<rightarrow> carrier G; g \<in> B \<rightarrow> carrier G\<rbrakk> |
|
394 |
\<Longrightarrow> finprod G g (A Un B) = finprod G g A \<otimes> finprod G g B" |
|
395 |
by (metis Pi_split_domain finprod_Un_Int finprod_closed finprod_empty r_one) |
|
13936 | 396 |
|
68517 | 397 |
lemma finprod_multf [simp]: |
68458 | 398 |
"\<lbrakk>f \<in> A \<rightarrow> carrier G; g \<in> A \<rightarrow> carrier G\<rbrakk> \<Longrightarrow> |
68517 | 399 |
finprod G (\<lambda>x. f x \<otimes> g x) A = (finprod G f A \<otimes> finprod G g A)" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
400 |
proof (induct A rule: infinite_finite_induct) |
13936 | 401 |
case empty show ?case by simp |
402 |
next |
|
15328 | 403 |
case (insert a A) then |
61384 | 404 |
have fA: "f \<in> A \<rightarrow> carrier G" by fast |
14750 | 405 |
from insert have fa: "f a \<in> carrier G" by fast |
61384 | 406 |
from insert have gA: "g \<in> A \<rightarrow> carrier G" by fast |
14750 | 407 |
from insert have ga: "g a \<in> carrier G" by fast |
61384 | 408 |
from insert have fgA: "(%x. f x \<otimes> g x) \<in> A \<rightarrow> carrier G" |
13936 | 409 |
by (simp add: Pi_def) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
410 |
show ?case |
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
411 |
by (simp add: insert fA fa gA ga fgA m_ac) |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
412 |
qed simp |
13936 | 413 |
|
27933 | 414 |
lemma finprod_cong': |
68458 | 415 |
"\<lbrakk>A = B; g \<in> B \<rightarrow> carrier G; |
416 |
!!i. i \<in> B \<Longrightarrow> f i = g i\<rbrakk> \<Longrightarrow> finprod G f A = finprod G g B" |
|
13936 | 417 |
proof - |
61384 | 418 |
assume prems: "A = B" "g \<in> B \<rightarrow> carrier G" |
68458 | 419 |
"!!i. i \<in> B \<Longrightarrow> f i = g i" |
13936 | 420 |
show ?thesis |
421 |
proof (cases "finite B") |
|
422 |
case True |
|
68458 | 423 |
then have "!!A. \<lbrakk>A = B; g \<in> B \<rightarrow> carrier G; |
424 |
!!i. i \<in> B \<Longrightarrow> f i = g i\<rbrakk> \<Longrightarrow> finprod G f A = finprod G g B" |
|
13936 | 425 |
proof induct |
426 |
case empty thus ?case by simp |
|
427 |
next |
|
15328 | 428 |
case (insert x B) |
13936 | 429 |
then have "finprod G f A = finprod G f (insert x B)" by simp |
430 |
also from insert have "... = f x \<otimes> finprod G f B" |
|
431 |
proof (intro finprod_insert) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
432 |
show "finite B" by fact |
13936 | 433 |
next |
67613 | 434 |
show "x \<notin> B" by fact |
13936 | 435 |
next |
67613 | 436 |
assume "x \<notin> B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
437 |
"g \<in> insert x B \<rightarrow> carrier G" |
61384 | 438 |
thus "f \<in> B \<rightarrow> carrier G" by fastforce |
13936 | 439 |
next |
67613 | 440 |
assume "x \<notin> B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
441 |
"g \<in> insert x B \<rightarrow> carrier G" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
442 |
thus "f x \<in> carrier G" by fastforce |
13936 | 443 |
qed |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
444 |
also from insert have "... = g x \<otimes> finprod G g B" by fastforce |
13936 | 445 |
also from insert have "... = finprod G g (insert x B)" |
446 |
by (intro finprod_insert [THEN sym]) auto |
|
447 |
finally show ?case . |
|
448 |
qed |
|
449 |
with prems show ?thesis by simp |
|
450 |
next |
|
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
451 |
case False with prems show ?thesis by simp |
13936 | 452 |
qed |
453 |
qed |
|
454 |
||
27933 | 455 |
lemma finprod_cong: |
68458 | 456 |
"\<lbrakk>A = B; f \<in> B \<rightarrow> carrier G = True; |
457 |
\<And>i. i \<in> B =simp=> f i = g i\<rbrakk> \<Longrightarrow> finprod G f A = finprod G g B" |
|
14213
7bf882b0a51e
Changed order of prems in finprod_cong. Slight speedup.
ballarin
parents:
13936
diff
changeset
|
458 |
(* This order of prems is slightly faster (3%) than the last two swapped. *) |
41433
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents:
40786
diff
changeset
|
459 |
by (rule finprod_cong') (auto simp add: simp_implies_def) |
13936 | 460 |
|
61382 | 461 |
text \<open>Usually, if this rule causes a failed congruence proof error, |
63167 | 462 |
the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown. |
13936 | 463 |
Adding @{thm [source] Pi_def} to the simpset is often useful. |
56142
8bb21318e10b
tuned -- command 'text' was localized some years ago;
wenzelm
parents:
46721
diff
changeset
|
464 |
For this reason, @{thm [source] finprod_cong} |
13936 | 465 |
is not added to the simpset by default. |
61382 | 466 |
\<close> |
13936 | 467 |
|
27933 | 468 |
end |
469 |
||
13936 | 470 |
declare funcsetI [rule del] |
471 |
funcset_mem [rule del] |
|
472 |
||
27933 | 473 |
context comm_monoid begin |
474 |
||
475 |
lemma finprod_0 [simp]: |
|
68458 | 476 |
"f \<in> {0::nat} \<rightarrow> carrier G \<Longrightarrow> finprod G f {..0} = f 0" |
68517 | 477 |
by (simp add: Pi_def) |
478 |
||
479 |
lemma finprod_0': |
|
480 |
"f \<in> {..n} \<rightarrow> carrier G \<Longrightarrow> (f 0) \<otimes> finprod G f {Suc 0..n} = finprod G f {..n}" |
|
481 |
proof - |
|
482 |
assume A: "f \<in> {.. n} \<rightarrow> carrier G" |
|
483 |
hence "(f 0) \<otimes> finprod G f {Suc 0..n} = finprod G f {..0} \<otimes> finprod G f {Suc 0..n}" |
|
484 |
using finprod_0[of f] by (simp add: funcset_mem) |
|
485 |
also have " ... = finprod G f ({..0} \<union> {Suc 0..n})" |
|
486 |
using finprod_Un_disjoint[of "{..0}" "{Suc 0..n}" f] A by (simp add: funcset_mem) |
|
487 |
also have " ... = finprod G f {..n}" |
|
488 |
by (simp add: atLeastAtMost_insertL atMost_atLeast0) |
|
489 |
finally show ?thesis . |
|
490 |
qed |
|
13936 | 491 |
|
27933 | 492 |
lemma finprod_Suc [simp]: |
68458 | 493 |
"f \<in> {..Suc n} \<rightarrow> carrier G \<Longrightarrow> |
13936 | 494 |
finprod G f {..Suc n} = (f (Suc n) \<otimes> finprod G f {..n})" |
495 |
by (simp add: Pi_def atMost_Suc) |
|
496 |
||
27933 | 497 |
lemma finprod_Suc2: |
68458 | 498 |
"f \<in> {..Suc n} \<rightarrow> carrier G \<Longrightarrow> |
13936 | 499 |
finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} \<otimes> f 0)" |
500 |
proof (induct n) |
|
501 |
case 0 thus ?case by (simp add: Pi_def) |
|
502 |
next |
|
503 |
case Suc thus ?case by (simp add: m_assoc Pi_def) |
|
504 |
qed |
|
505 |
||
68517 | 506 |
lemma finprod_Suc3: |
507 |
assumes "f \<in> {..n :: nat} \<rightarrow> carrier G" |
|
508 |
shows "finprod G f {.. n} = (f n) \<otimes> finprod G f {..< n}" |
|
509 |
proof (cases "n = 0") |
|
510 |
case True thus ?thesis |
|
511 |
using assms atMost_Suc by simp |
|
512 |
next |
|
513 |
case False |
|
514 |
then obtain k where "n = Suc k" |
|
515 |
using not0_implies_Suc by blast |
|
516 |
thus ?thesis |
|
517 |
using finprod_Suc[of f k] assms atMost_Suc lessThan_Suc_atMost by simp |
|
518 |
qed |
|
13936 | 519 |
|
69895
6b03a8cf092d
more formal contributors (with the help of the history);
wenzelm
parents:
69313
diff
changeset
|
520 |
lemma finprod_reindex: \<^marker>\<open>contributor \<open>Jeremy Avigad\<close>\<close> |
68458 | 521 |
"f \<in> (h ` A) \<rightarrow> carrier G \<Longrightarrow> |
67613 | 522 |
inj_on h A \<Longrightarrow> finprod G f (h ` A) = finprod G (\<lambda>x. f (h x)) A" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
523 |
proof (induct A rule: infinite_finite_induct) |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
524 |
case (infinite A) |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
525 |
hence "\<not> finite (h ` A)" |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
526 |
using finite_imageD by blast |
61382 | 527 |
with \<open>\<not> finite A\<close> show ?case by simp |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
528 |
qed (auto simp add: Pi_def) |
27699 | 529 |
|
69895
6b03a8cf092d
more formal contributors (with the help of the history);
wenzelm
parents:
69313
diff
changeset
|
530 |
lemma finprod_const: \<^marker>\<open>contributor \<open>Jeremy Avigad\<close>\<close> |
67613 | 531 |
assumes a [simp]: "a \<in> carrier G" |
532 |
shows "finprod G (\<lambda>x. a) A = a [^] card A" |
|
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
533 |
proof (induct A rule: infinite_finite_induct) |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
534 |
case (insert b A) |
68458 | 535 |
show ?case |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
536 |
proof (subst finprod_insert[OF insert(1-2)]) |
67341
df79ef3b3a41
Renamed (^) to [^] in preparation of the move from "op X" to (X)
nipkow
parents:
67091
diff
changeset
|
537 |
show "a \<otimes> (\<Otimes>x\<in>A. a) = a [^] card (insert b A)" |
60112
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
538 |
by (insert insert, auto, subst m_comm, auto) |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
539 |
qed auto |
3eab4acaa035
finprod takes 1 in case of infinite sets => remove several "finite A" assumptions
Rene Thiemann <rene.thiemann@uibk.ac.at>
parents:
58860
diff
changeset
|
540 |
qed auto |
27699 | 541 |
|
69895
6b03a8cf092d
more formal contributors (with the help of the history);
wenzelm
parents:
69313
diff
changeset
|
542 |
lemma finprod_singleton: \<^marker>\<open>contributor \<open>Jesus Aransay\<close>\<close> |
27933 | 543 |
assumes i_in_A: "i \<in> A" and fin_A: "finite A" and f_Pi: "f \<in> A \<rightarrow> carrier G" |
544 |
shows "(\<Otimes>j\<in>A. if i = j then f j else \<one>) = f i" |
|
29237 | 545 |
using i_in_A finprod_insert [of "A - {i}" i "(\<lambda>j. if i = j then f j else \<one>)"] |
546 |
fin_A f_Pi finprod_one [of "A - {i}"] |
|
68458 | 547 |
finprod_cong [of "A - {i}" "A - {i}" "(\<lambda>j. if i = j then f j else \<one>)" "(\<lambda>i. \<one>)"] |
27933 | 548 |
unfolding Pi_def simp_implies_def by (force simp add: insert_absorb) |
549 |
||
70044
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
550 |
lemma finprod_singleton_swap: |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
551 |
assumes i_in_A: "i \<in> A" and fin_A: "finite A" and f_Pi: "f \<in> A \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
552 |
shows "(\<Otimes>j\<in>A. if j = i then f j else \<one>) = f i" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
553 |
using finprod_singleton [OF assms] by (simp add: eq_commute) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
554 |
|
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
555 |
lemma finprod_mono_neutral_cong_left: |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
556 |
assumes "finite B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
557 |
and "A \<subseteq> B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
558 |
and 1: "\<And>i. i \<in> B - A \<Longrightarrow> h i = \<one>" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
559 |
and gh: "\<And>x. x \<in> A \<Longrightarrow> g x = h x" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
560 |
and h: "h \<in> B \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
561 |
shows "finprod G g A = finprod G h B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
562 |
proof- |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
563 |
have eq: "A \<union> (B - A) = B" using \<open>A \<subseteq> B\<close> by blast |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
564 |
have d: "A \<inter> (B - A) = {}" using \<open>A \<subseteq> B\<close> by blast |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
565 |
from \<open>finite B\<close> \<open>A \<subseteq> B\<close> have f: "finite A" "finite (B - A)" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
566 |
by (auto intro: finite_subset) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
567 |
have "h \<in> A \<rightarrow> carrier G" "h \<in> B - A \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
568 |
using assms by (auto simp: image_subset_iff_funcset) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
569 |
moreover have "finprod G g A = finprod G h A \<otimes> finprod G h (B - A)" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
570 |
proof - |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
571 |
have "finprod G h (B - A) = \<one>" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
572 |
using "1" finprod_one_eqI by blast |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
573 |
moreover have "finprod G g A = finprod G h A" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
574 |
using \<open>h \<in> A \<rightarrow> carrier G\<close> finprod_cong' gh by blast |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
575 |
ultimately show ?thesis |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
576 |
by (simp add: \<open>h \<in> A \<rightarrow> carrier G\<close>) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
577 |
qed |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
578 |
ultimately show ?thesis |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
579 |
by (simp add: finprod_Un_disjoint [OF f d, unfolded eq]) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
580 |
qed |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
581 |
|
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
582 |
lemma finprod_mono_neutral_cong_right: |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
583 |
assumes "finite B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
584 |
and "A \<subseteq> B" "\<And>i. i \<in> B - A \<Longrightarrow> g i = \<one>" "\<And>x. x \<in> A \<Longrightarrow> g x = h x" "g \<in> B \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
585 |
shows "finprod G g B = finprod G h A" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
586 |
using assms by (auto intro!: finprod_mono_neutral_cong_left [symmetric]) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
587 |
|
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
588 |
lemma finprod_mono_neutral_cong: |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
589 |
assumes [simp]: "finite B" "finite A" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
590 |
and *: "\<And>i. i \<in> B - A \<Longrightarrow> h i = \<one>" "\<And>i. i \<in> A - B \<Longrightarrow> g i = \<one>" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
591 |
and gh: "\<And>x. x \<in> A \<inter> B \<Longrightarrow> g x = h x" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
592 |
and g: "g \<in> A \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
593 |
and h: "h \<in> B \<rightarrow> carrier G" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
594 |
shows "finprod G g A = finprod G h B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
595 |
proof- |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
596 |
have "finprod G g A = finprod G g (A \<inter> B)" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
597 |
by (rule finprod_mono_neutral_cong_right) (use assms in auto) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
598 |
also have "\<dots> = finprod G h (A \<inter> B)" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
599 |
by (rule finprod_cong) (use assms in auto) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
600 |
also have "\<dots> = finprod G h B" |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
601 |
by (rule finprod_mono_neutral_cong_left) (use assms in auto) |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
602 |
finally show ?thesis . |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
603 |
qed |
da5857dbcbb9
More group theory. Sum and product indexed by the non-neutral part of a set
paulson <lp15@cam.ac.uk>
parents:
69895
diff
changeset
|
604 |
|
13936 | 605 |
end |
27933 | 606 |
|
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
607 |
(* Jeremy Avigad. This should be generalized to arbitrary groups, not just commutative |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
608 |
ones, using Lagrange's theorem. *) |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
609 |
|
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
610 |
lemma (in comm_group) power_order_eq_one: |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
611 |
assumes fin [simp]: "finite (carrier G)" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
612 |
and a [simp]: "a \<in> carrier G" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
613 |
shows "a [^] card(carrier G) = one G" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
614 |
proof - |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
615 |
have "(\<Otimes>x\<in>carrier G. x) = (\<Otimes>x\<in>carrier G. a \<otimes> x)" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
616 |
by (subst (2) finprod_reindex [symmetric], |
68458 | 617 |
auto simp add: Pi_def inj_on_cmult surj_const_mult) |
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
618 |
also have "\<dots> = (\<Otimes>x\<in>carrier G. a) \<otimes> (\<Otimes>x\<in>carrier G. x)" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
619 |
by (auto simp add: finprod_multf Pi_def) |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
620 |
also have "(\<Otimes>x\<in>carrier G. a) = a [^] card(carrier G)" |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
621 |
by (auto simp add: finprod_const) |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
622 |
finally show ?thesis |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
623 |
by auto |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
624 |
qed |
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
625 |
|
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
626 |
lemma (in comm_monoid) finprod_UN_disjoint: |
68458 | 627 |
assumes |
628 |
"finite I" "\<And>i. i \<in> I \<Longrightarrow> finite (A i)" "pairwise (\<lambda>i j. disjnt (A i) (A j)) I" |
|
629 |
"\<And>i x. i \<in> I \<Longrightarrow> x \<in> A i \<Longrightarrow> g x \<in> carrier G" |
|
69313 | 630 |
shows "finprod G g (\<Union>(A ` I)) = finprod G (\<lambda>i. finprod G g (A i)) I" |
68458 | 631 |
using assms |
632 |
proof (induction set: finite) |
|
633 |
case empty |
|
634 |
then show ?case |
|
635 |
by force |
|
636 |
next |
|
637 |
case (insert i I) |
|
638 |
then show ?case |
|
639 |
unfolding pairwise_def disjnt_def |
|
640 |
apply clarsimp |
|
641 |
apply (subst finprod_Un_disjoint) |
|
642 |
apply (fastforce intro!: funcsetI finprod_closed)+ |
|
643 |
done |
|
644 |
qed |
|
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
645 |
|
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
646 |
lemma (in comm_monoid) finprod_Union_disjoint: |
68458 | 647 |
"\<lbrakk>finite C; \<And>A. A \<in> C \<Longrightarrow> finite A \<and> (\<forall>x\<in>A. f x \<in> carrier G); pairwise disjnt C\<rbrakk> \<Longrightarrow> |
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
648 |
finprod G f (\<Union>C) = finprod G (finprod G f) C" |
68517 | 649 |
by (frule finprod_UN_disjoint [of C id f]) auto |
68445
c183a6a69f2d
reorganisation of Algebra: new material from Baillon and Vilhena, removal of duplicate names, elimination of "More_" theories
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
650 |
|
27933 | 651 |
end |