author | haftmann |
Fri, 02 Mar 2012 21:45:45 +0100 | |
changeset 46767 | 807a5d219c23 |
parent 46721 | f88b187ad8ca |
child 56142 | 8bb21318e10b |
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 |
|
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27699
diff
changeset
|
11 |
subsection {* Product Operator for Commutative Monoids *} |
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
12 |
|
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27699
diff
changeset
|
13 |
subsubsection {* Inductive Definition of a Relation for Products over Sets *} |
13936 | 14 |
|
14750 | 15 |
text {* Instantiation of locale @{text LC} of theory @{text Finite_Set} is not |
16 |
possible, because here we have explicit typing rules like |
|
17 |
@{text "x \<in> carrier G"}. We introduce an explicit argument for the domain |
|
14651 | 18 |
@{text D}. *} |
13936 | 19 |
|
23746 | 20 |
inductive_set |
13936 | 21 |
foldSetD :: "['a set, 'b => 'a => 'a, 'a] => ('b set * 'a) set" |
23746 | 22 |
for D :: "'a set" and f :: "'b => 'a => 'a" and e :: 'a |
23 |
where |
|
14750 | 24 |
emptyI [intro]: "e \<in> D ==> ({}, e) \<in> foldSetD D f e" |
23746 | 25 |
| insertI [intro]: "[| x ~: A; f x y \<in> D; (A, y) \<in> foldSetD D f e |] ==> |
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 |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
31 |
foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a" |
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 |
|
34 |
lemma foldSetD_closed: |
|
14750 | 35 |
"[| (A, z) \<in> foldSetD D f e ; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D |
36 |
|] ==> z \<in> D"; |
|
23746 | 37 |
by (erule foldSetD.cases) auto |
13936 | 38 |
|
39 |
lemma Diff1_foldSetD: |
|
14750 | 40 |
"[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; f x y \<in> D |] ==> |
41 |
(A, f x y) \<in> foldSetD D f e" |
|
13936 | 42 |
apply (erule insert_Diff [THEN subst], rule foldSetD.intros) |
43 |
apply auto |
|
44 |
done |
|
45 |
||
14750 | 46 |
lemma foldSetD_imp_finite [simp]: "(A, x) \<in> foldSetD D f e ==> finite A" |
13936 | 47 |
by (induct set: foldSetD) auto |
48 |
||
49 |
lemma finite_imp_foldSetD: |
|
14750 | 50 |
"[| finite A; e \<in> D; !!x y. [| x \<in> A; y \<in> D |] ==> f x y \<in> D |] ==> |
51 |
EX x. (A, x) \<in> foldSetD D f e" |
|
22265 | 52 |
proof (induct set: finite) |
13936 | 53 |
case empty then show ?case by auto |
54 |
next |
|
15328 | 55 |
case (insert x F) |
14750 | 56 |
then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto |
57 |
with insert have "y \<in> D" by (auto dest: foldSetD_closed) |
|
58 |
with y and insert have "(insert x F, f x y) \<in> foldSetD D f e" |
|
13936 | 59 |
by (intro foldSetD.intros) auto |
60 |
then show ?case .. |
|
61 |
qed |
|
62 |
||
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
63 |
|
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27699
diff
changeset
|
64 |
text {* Left-Commutative Operations *} |
13936 | 65 |
|
66 |
locale LCD = |
|
67 |
fixes B :: "'b set" |
|
68 |
and D :: "'a set" |
|
69 |
and f :: "'b => 'a => 'a" (infixl "\<cdot>" 70) |
|
70 |
assumes left_commute: |
|
14750 | 71 |
"[| x \<in> B; y \<in> B; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
72 |
and f_closed [simp, intro!]: "!!x y. [| x \<in> B; y \<in> D |] ==> f x y \<in> D" |
|
13936 | 73 |
|
74 |
lemma (in LCD) foldSetD_closed [dest]: |
|
14750 | 75 |
"(A, z) \<in> foldSetD D f e ==> z \<in> D"; |
23746 | 76 |
by (erule foldSetD.cases) auto |
13936 | 77 |
|
78 |
lemma (in LCD) Diff1_foldSetD: |
|
14750 | 79 |
"[| (A - {x}, y) \<in> foldSetD D f e; x \<in> A; A \<subseteq> B |] ==> |
80 |
(A, f x y) \<in> foldSetD D f e" |
|
81 |
apply (subgoal_tac "x \<in> B") |
|
13936 | 82 |
prefer 2 apply fast |
83 |
apply (erule insert_Diff [THEN subst], rule foldSetD.intros) |
|
84 |
apply auto |
|
85 |
done |
|
86 |
||
87 |
lemma (in LCD) foldSetD_imp_finite [simp]: |
|
14750 | 88 |
"(A, x) \<in> foldSetD D f e ==> finite A" |
13936 | 89 |
by (induct set: foldSetD) auto |
90 |
||
91 |
lemma (in LCD) finite_imp_foldSetD: |
|
14750 | 92 |
"[| finite A; A \<subseteq> B; e \<in> D |] ==> EX x. (A, x) \<in> foldSetD D f e" |
22265 | 93 |
proof (induct set: finite) |
13936 | 94 |
case empty then show ?case by auto |
95 |
next |
|
15328 | 96 |
case (insert x F) |
14750 | 97 |
then obtain y where y: "(F, y) \<in> foldSetD D f e" by auto |
98 |
with insert have "y \<in> D" by auto |
|
99 |
with y and insert have "(insert x F, f x y) \<in> foldSetD D f e" |
|
13936 | 100 |
by (intro foldSetD.intros) auto |
101 |
then show ?case .. |
|
102 |
qed |
|
103 |
||
104 |
lemma (in LCD) foldSetD_determ_aux: |
|
14750 | 105 |
"e \<in> D ==> \<forall>A x. A \<subseteq> B & card A < n --> (A, x) \<in> foldSetD D f e --> |
106 |
(\<forall>y. (A, y) \<in> foldSetD D f e --> y = x)" |
|
13936 | 107 |
apply (induct n) |
108 |
apply (auto simp add: less_Suc_eq) (* slow *) |
|
109 |
apply (erule foldSetD.cases) |
|
110 |
apply blast |
|
111 |
apply (erule foldSetD.cases) |
|
112 |
apply blast |
|
113 |
apply clarify |
|
114 |
txt {* force simplification of @{text "card A < card (insert ...)"}. *} |
|
115 |
apply (erule rev_mp) |
|
116 |
apply (simp add: less_Suc_eq_le) |
|
117 |
apply (rule impI) |
|
23746 | 118 |
apply (rename_tac xa Aa ya xb Ab yb, case_tac "xa = xb") |
13936 | 119 |
apply (subgoal_tac "Aa = Ab") |
120 |
prefer 2 apply (blast elim!: equalityE) |
|
121 |
apply blast |
|
122 |
txt {* case @{prop "xa \<notin> xb"}. *} |
|
14750 | 123 |
apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb \<in> Aa & xa \<in> Ab") |
13936 | 124 |
prefer 2 apply (blast elim!: equalityE) |
125 |
apply clarify |
|
126 |
apply (subgoal_tac "Aa = insert xb Ab - {xa}") |
|
127 |
prefer 2 apply blast |
|
14750 | 128 |
apply (subgoal_tac "card Aa \<le> card Ab") |
13936 | 129 |
prefer 2 |
130 |
apply (rule Suc_le_mono [THEN subst]) |
|
131 |
apply (simp add: card_Suc_Diff1) |
|
132 |
apply (rule_tac A1 = "Aa - {xb}" in finite_imp_foldSetD [THEN exE]) |
|
40786
0a54cfc9add3
gave more standard finite set rules simp and intro attribute
nipkow
parents:
35849
diff
changeset
|
133 |
apply (blast intro: foldSetD_imp_finite) |
13936 | 134 |
apply best |
135 |
apply assumption |
|
136 |
apply (frule (1) Diff1_foldSetD) |
|
137 |
apply best |
|
138 |
apply (subgoal_tac "ya = f xb x") |
|
139 |
prefer 2 |
|
14750 | 140 |
apply (subgoal_tac "Aa \<subseteq> B") |
13936 | 141 |
prefer 2 apply best (* slow *) |
142 |
apply (blast del: equalityCE) |
|
14750 | 143 |
apply (subgoal_tac "(Ab - {xa}, x) \<in> foldSetD D f e") |
13936 | 144 |
prefer 2 apply simp |
145 |
apply (subgoal_tac "yb = f xa x") |
|
146 |
prefer 2 |
|
147 |
apply (blast del: equalityCE dest: Diff1_foldSetD) |
|
148 |
apply (simp (no_asm_simp)) |
|
149 |
apply (rule left_commute) |
|
150 |
apply assumption |
|
151 |
apply best (* slow *) |
|
152 |
apply best |
|
153 |
done |
|
154 |
||
155 |
lemma (in LCD) foldSetD_determ: |
|
14750 | 156 |
"[| (A, x) \<in> foldSetD D f e; (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |] |
13936 | 157 |
==> y = x" |
158 |
by (blast intro: foldSetD_determ_aux [rule_format]) |
|
159 |
||
160 |
lemma (in LCD) foldD_equality: |
|
14750 | 161 |
"[| (A, y) \<in> foldSetD D f e; e \<in> D; A \<subseteq> B |] ==> foldD D f e A = y" |
13936 | 162 |
by (unfold foldD_def) (blast intro: foldSetD_determ) |
163 |
||
164 |
lemma foldD_empty [simp]: |
|
14750 | 165 |
"e \<in> D ==> foldD D f e {} = e" |
13936 | 166 |
by (unfold foldD_def) blast |
167 |
||
168 |
lemma (in LCD) foldD_insert_aux: |
|
14750 | 169 |
"[| x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==> |
170 |
((insert x A, v) \<in> foldSetD D f e) = |
|
171 |
(EX y. (A, y) \<in> foldSetD D f e & v = f x y)" |
|
13936 | 172 |
apply auto |
173 |
apply (rule_tac A1 = A in finite_imp_foldSetD [THEN exE]) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
174 |
apply (fastforce dest: foldSetD_imp_finite) |
13936 | 175 |
apply assumption |
176 |
apply assumption |
|
177 |
apply (blast intro: foldSetD_determ) |
|
178 |
done |
|
179 |
||
180 |
lemma (in LCD) foldD_insert: |
|
14750 | 181 |
"[| finite A; x ~: A; x \<in> B; e \<in> D; A \<subseteq> B |] ==> |
13936 | 182 |
foldD D f e (insert x A) = f x (foldD D f e A)" |
183 |
apply (unfold foldD_def) |
|
184 |
apply (simp add: foldD_insert_aux) |
|
185 |
apply (rule the_equality) |
|
186 |
apply (auto intro: finite_imp_foldSetD |
|
187 |
cong add: conj_cong simp add: foldD_def [symmetric] foldD_equality) |
|
188 |
done |
|
189 |
||
190 |
lemma (in LCD) foldD_closed [simp]: |
|
14750 | 191 |
"[| finite A; e \<in> D; A \<subseteq> B |] ==> foldD D f e A \<in> D" |
22265 | 192 |
proof (induct set: finite) |
46721 | 193 |
case empty then show ?case by simp |
13936 | 194 |
next |
195 |
case insert then show ?case by (simp add: foldD_insert) |
|
196 |
qed |
|
197 |
||
198 |
lemma (in LCD) foldD_commute: |
|
14750 | 199 |
"[| finite A; x \<in> B; e \<in> D; A \<subseteq> B |] ==> |
13936 | 200 |
f x (foldD D f e A) = foldD D f (f x e) A" |
22265 | 201 |
apply (induct set: finite) |
13936 | 202 |
apply simp |
203 |
apply (auto simp add: left_commute foldD_insert) |
|
204 |
done |
|
205 |
||
206 |
lemma Int_mono2: |
|
14750 | 207 |
"[| A \<subseteq> C; B \<subseteq> C |] ==> A Int B \<subseteq> C" |
13936 | 208 |
by blast |
209 |
||
210 |
lemma (in LCD) foldD_nest_Un_Int: |
|
14750 | 211 |
"[| finite A; finite C; e \<in> D; A \<subseteq> B; C \<subseteq> B |] ==> |
13936 | 212 |
foldD D f (foldD D f e C) A = foldD D f (foldD D f e (A Int C)) (A Un C)" |
22265 | 213 |
apply (induct set: finite) |
13936 | 214 |
apply simp |
215 |
apply (simp add: foldD_insert foldD_commute Int_insert_left insert_absorb |
|
32693 | 216 |
Int_mono2) |
13936 | 217 |
done |
218 |
||
219 |
lemma (in LCD) foldD_nest_Un_disjoint: |
|
14750 | 220 |
"[| finite A; finite B; A Int B = {}; e \<in> D; A \<subseteq> B; C \<subseteq> B |] |
13936 | 221 |
==> foldD D f e (A Un B) = foldD D f (foldD D f e B) A" |
222 |
by (simp add: foldD_nest_Un_Int) |
|
223 |
||
224 |
-- {* Delete rules to do with @{text foldSetD} relation. *} |
|
225 |
||
226 |
declare foldSetD_imp_finite [simp del] |
|
227 |
empty_foldSetDE [rule del] |
|
228 |
foldSetD.intros [rule del] |
|
229 |
declare (in LCD) |
|
230 |
foldSetD_closed [rule del] |
|
231 |
||
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
232 |
|
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27699
diff
changeset
|
233 |
text {* Commutative Monoids *} |
13936 | 234 |
|
235 |
text {* |
|
236 |
We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"} |
|
237 |
instead of @{text "'b => 'a => 'a"}. |
|
238 |
*} |
|
239 |
||
240 |
locale ACeD = |
|
241 |
fixes D :: "'a set" |
|
242 |
and f :: "'a => 'a => 'a" (infixl "\<cdot>" 70) |
|
243 |
and e :: 'a |
|
14750 | 244 |
assumes ident [simp]: "x \<in> D ==> x \<cdot> e = x" |
245 |
and commute: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y = y \<cdot> x" |
|
246 |
and assoc: "[| x \<in> D; y \<in> D; z \<in> D |] ==> (x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
247 |
and e_closed [simp]: "e \<in> D" |
|
248 |
and f_closed [simp]: "[| x \<in> D; y \<in> D |] ==> x \<cdot> y \<in> D" |
|
13936 | 249 |
|
250 |
lemma (in ACeD) left_commute: |
|
14750 | 251 |
"[| x \<in> D; y \<in> D; z \<in> D |] ==> x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
13936 | 252 |
proof - |
14750 | 253 |
assume D: "x \<in> D" "y \<in> D" "z \<in> D" |
13936 | 254 |
then have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp add: commute) |
255 |
also from D have "... = y \<cdot> (z \<cdot> x)" by (simp add: assoc) |
|
256 |
also from D have "z \<cdot> x = x \<cdot> z" by (simp add: commute) |
|
257 |
finally show ?thesis . |
|
258 |
qed |
|
259 |
||
260 |
lemmas (in ACeD) AC = assoc commute left_commute |
|
261 |
||
14750 | 262 |
lemma (in ACeD) left_ident [simp]: "x \<in> D ==> e \<cdot> x = x" |
13936 | 263 |
proof - |
23350 | 264 |
assume "x \<in> D" |
265 |
then have "x \<cdot> e = x" by (rule ident) |
|
266 |
with `x \<in> D` show ?thesis by (simp add: commute) |
|
13936 | 267 |
qed |
268 |
||
269 |
lemma (in ACeD) foldD_Un_Int: |
|
14750 | 270 |
"[| finite A; finite B; A \<subseteq> D; B \<subseteq> D |] ==> |
13936 | 271 |
foldD D f e A \<cdot> foldD D f e B = |
272 |
foldD D f e (A Un B) \<cdot> foldD D f e (A Int B)" |
|
22265 | 273 |
apply (induct set: finite) |
13936 | 274 |
apply (simp add: left_commute LCD.foldD_closed [OF LCD.intro [of D]]) |
275 |
apply (simp add: AC insert_absorb Int_insert_left |
|
276 |
LCD.foldD_insert [OF LCD.intro [of D]] |
|
277 |
LCD.foldD_closed [OF LCD.intro [of D]] |
|
32693 | 278 |
Int_mono2) |
13936 | 279 |
done |
280 |
||
281 |
lemma (in ACeD) foldD_Un_disjoint: |
|
14750 | 282 |
"[| finite A; finite B; A Int B = {}; A \<subseteq> D; B \<subseteq> D |] ==> |
13936 | 283 |
foldD D f e (A Un B) = foldD D f e A \<cdot> foldD D f e B" |
284 |
by (simp add: foldD_Un_Int |
|
32693 | 285 |
left_commute LCD.foldD_closed [OF LCD.intro [of D]]) |
13936 | 286 |
|
20318
0e0ea63fe768
Restructured algebra library, added ideals and quotient rings.
ballarin
parents:
16638
diff
changeset
|
287 |
|
27717
21bbd410ba04
Generalised polynomial lemmas from cring to ring.
ballarin
parents:
27699
diff
changeset
|
288 |
subsubsection {* Products over Finite Sets *} |
13936 | 289 |
|
35847 | 290 |
definition |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
291 |
finprod :: "[('b, 'm) monoid_scheme, 'a => 'b, 'a set] => 'b" |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
292 |
where "finprod G f A = |
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
293 |
(if finite A |
35847 | 294 |
then foldD (carrier G) (mult G o f) \<one>\<^bsub>G\<^esub> A |
35848
5443079512ea
slightly more uniform definitions -- eliminated old-style meta-equality;
wenzelm
parents:
35847
diff
changeset
|
295 |
else undefined)" |
13936 | 296 |
|
14651 | 297 |
syntax |
298 |
"_finprod" :: "index => idt => 'a set => 'b => 'b" |
|
14666 | 299 |
("(3\<Otimes>__:_. _)" [1000, 0, 51, 10] 10) |
14651 | 300 |
syntax (xsymbols) |
301 |
"_finprod" :: "index => idt => 'a set => 'b => 'b" |
|
14666 | 302 |
("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10) |
14651 | 303 |
syntax (HTML output) |
304 |
"_finprod" :: "index => idt => 'a set => 'b => 'b" |
|
14666 | 305 |
("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10) |
14651 | 306 |
translations |
35054 | 307 |
"\<Otimes>\<index>i:A. b" == "CONST finprod \<struct>\<index> (%i. b) A" |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
308 |
-- {* Beware of argument permutation! *} |
13936 | 309 |
|
310 |
lemma (in comm_monoid) finprod_empty [simp]: |
|
311 |
"finprod G f {} = \<one>" |
|
312 |
by (simp add: finprod_def) |
|
313 |
||
314 |
declare funcsetI [intro] |
|
315 |
funcset_mem [dest] |
|
316 |
||
27933 | 317 |
context comm_monoid begin |
318 |
||
319 |
lemma finprod_insert [simp]: |
|
13936 | 320 |
"[| finite F; a \<notin> F; f \<in> F -> carrier G; f a \<in> carrier G |] ==> |
321 |
finprod G f (insert a F) = f a \<otimes> finprod G f F" |
|
322 |
apply (rule trans) |
|
323 |
apply (simp add: finprod_def) |
|
324 |
apply (rule trans) |
|
325 |
apply (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]]) |
|
326 |
apply simp |
|
327 |
apply (rule m_lcomm) |
|
328 |
apply fast |
|
329 |
apply fast |
|
330 |
apply assumption |
|
46721 | 331 |
apply fastforce |
13936 | 332 |
apply simp+ |
333 |
apply fast |
|
334 |
apply (auto simp add: finprod_def) |
|
335 |
done |
|
336 |
||
27933 | 337 |
lemma finprod_one [simp]: |
14651 | 338 |
"finite A ==> (\<Otimes>i:A. \<one>) = \<one>" |
22265 | 339 |
proof (induct set: finite) |
13936 | 340 |
case empty show ?case by simp |
341 |
next |
|
15328 | 342 |
case (insert a A) |
13936 | 343 |
have "(%i. \<one>) \<in> A -> carrier G" by auto |
344 |
with insert show ?case by simp |
|
345 |
qed |
|
346 |
||
27933 | 347 |
lemma finprod_closed [simp]: |
13936 | 348 |
fixes A |
349 |
assumes fin: "finite A" and f: "f \<in> A -> carrier G" |
|
350 |
shows "finprod G f A \<in> carrier G" |
|
351 |
using fin f |
|
352 |
proof induct |
|
353 |
case empty show ?case by simp |
|
354 |
next |
|
15328 | 355 |
case (insert a A) |
13936 | 356 |
then have a: "f a \<in> carrier G" by fast |
357 |
from insert have A: "f \<in> A -> carrier G" by fast |
|
358 |
from insert A a show ?case by simp |
|
359 |
qed |
|
360 |
||
361 |
lemma funcset_Int_left [simp, intro]: |
|
362 |
"[| f \<in> A -> C; f \<in> B -> C |] ==> f \<in> A Int B -> C" |
|
363 |
by fast |
|
364 |
||
365 |
lemma funcset_Un_left [iff]: |
|
366 |
"(f \<in> A Un B -> C) = (f \<in> A -> C & f \<in> B -> C)" |
|
367 |
by fast |
|
368 |
||
27933 | 369 |
lemma finprod_Un_Int: |
13936 | 370 |
"[| finite A; finite B; g \<in> A -> carrier G; g \<in> B -> carrier G |] ==> |
371 |
finprod G g (A Un B) \<otimes> finprod G g (A Int B) = |
|
372 |
finprod G g A \<otimes> finprod G g B" |
|
373 |
-- {* The reversed orientation looks more natural, but LOOPS as a simprule! *} |
|
22265 | 374 |
proof (induct set: finite) |
46721 | 375 |
case empty then show ?case by simp |
13936 | 376 |
next |
15328 | 377 |
case (insert a A) |
13936 | 378 |
then have a: "g a \<in> carrier G" by fast |
379 |
from insert have A: "g \<in> A -> carrier G" by fast |
|
380 |
from insert A a show ?case |
|
46721 | 381 |
by (simp add: m_ac Int_insert_left insert_absorb Int_mono2) |
13936 | 382 |
qed |
383 |
||
27933 | 384 |
lemma finprod_Un_disjoint: |
13936 | 385 |
"[| finite A; finite B; A Int B = {}; |
386 |
g \<in> A -> carrier G; g \<in> B -> carrier G |] |
|
387 |
==> finprod G g (A Un B) = finprod G g A \<otimes> finprod G g B" |
|
388 |
apply (subst finprod_Un_Int [symmetric]) |
|
46721 | 389 |
apply auto |
13936 | 390 |
done |
391 |
||
27933 | 392 |
lemma finprod_multf: |
13936 | 393 |
"[| finite A; f \<in> A -> carrier G; g \<in> A -> carrier G |] ==> |
394 |
finprod G (%x. f x \<otimes> g x) A = (finprod G f A \<otimes> finprod G g A)" |
|
22265 | 395 |
proof (induct set: finite) |
13936 | 396 |
case empty show ?case by simp |
397 |
next |
|
15328 | 398 |
case (insert a A) then |
14750 | 399 |
have fA: "f \<in> A -> carrier G" by fast |
400 |
from insert have fa: "f a \<in> carrier G" by fast |
|
401 |
from insert have gA: "g \<in> A -> carrier G" by fast |
|
402 |
from insert have ga: "g a \<in> carrier G" by fast |
|
403 |
from insert have fgA: "(%x. f x \<otimes> g x) \<in> A -> carrier G" |
|
13936 | 404 |
by (simp add: Pi_def) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
405 |
show ?case |
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
406 |
by (simp add: insert fA fa gA ga fgA m_ac) |
13936 | 407 |
qed |
408 |
||
27933 | 409 |
lemma finprod_cong': |
14750 | 410 |
"[| A = B; g \<in> B -> carrier G; |
411 |
!!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B" |
|
13936 | 412 |
proof - |
14750 | 413 |
assume prems: "A = B" "g \<in> B -> carrier G" |
414 |
"!!i. i \<in> B ==> f i = g i" |
|
13936 | 415 |
show ?thesis |
416 |
proof (cases "finite B") |
|
417 |
case True |
|
14750 | 418 |
then have "!!A. [| A = B; g \<in> B -> carrier G; |
419 |
!!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B" |
|
13936 | 420 |
proof induct |
421 |
case empty thus ?case by simp |
|
422 |
next |
|
15328 | 423 |
case (insert x B) |
13936 | 424 |
then have "finprod G f A = finprod G f (insert x B)" by simp |
425 |
also from insert have "... = f x \<otimes> finprod G f B" |
|
426 |
proof (intro finprod_insert) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
427 |
show "finite B" by fact |
13936 | 428 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
429 |
show "x ~: B" by fact |
13936 | 430 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
431 |
assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
432 |
"g \<in> insert x B \<rightarrow> carrier G" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
433 |
thus "f \<in> B -> carrier G" by fastforce |
13936 | 434 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
435 |
assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
436 |
"g \<in> insert x B \<rightarrow> carrier G" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
437 |
thus "f x \<in> carrier G" by fastforce |
13936 | 438 |
qed |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
439 |
also from insert have "... = g x \<otimes> finprod G g B" by fastforce |
13936 | 440 |
also from insert have "... = finprod G g (insert x B)" |
441 |
by (intro finprod_insert [THEN sym]) auto |
|
442 |
finally show ?case . |
|
443 |
qed |
|
444 |
with prems show ?thesis by simp |
|
445 |
next |
|
446 |
case False with prems show ?thesis by (simp add: finprod_def) |
|
447 |
qed |
|
448 |
qed |
|
449 |
||
27933 | 450 |
lemma finprod_cong: |
14750 | 451 |
"[| A = B; f \<in> B -> carrier G = True; |
41433
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents:
40786
diff
changeset
|
452 |
!!i. i \<in> B =simp=> f i = g i |] ==> finprod G f A = finprod G g B" |
14213
7bf882b0a51e
Changed order of prems in finprod_cong. Slight speedup.
ballarin
parents:
13936
diff
changeset
|
453 |
(* 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
|
454 |
by (rule finprod_cong') (auto simp add: simp_implies_def) |
13936 | 455 |
|
456 |
text {*Usually, if this rule causes a failed congruence proof error, |
|
14750 | 457 |
the reason is that the premise @{text "g \<in> B -> carrier G"} cannot be shown. |
13936 | 458 |
Adding @{thm [source] Pi_def} to the simpset is often useful. |
459 |
For this reason, @{thm [source] comm_monoid.finprod_cong} |
|
460 |
is not added to the simpset by default. |
|
461 |
*} |
|
462 |
||
27933 | 463 |
end |
464 |
||
13936 | 465 |
declare funcsetI [rule del] |
466 |
funcset_mem [rule del] |
|
467 |
||
27933 | 468 |
context comm_monoid begin |
469 |
||
470 |
lemma finprod_0 [simp]: |
|
14750 | 471 |
"f \<in> {0::nat} -> carrier G ==> finprod G f {..0} = f 0" |
13936 | 472 |
by (simp add: Pi_def) |
473 |
||
27933 | 474 |
lemma finprod_Suc [simp]: |
14750 | 475 |
"f \<in> {..Suc n} -> carrier G ==> |
13936 | 476 |
finprod G f {..Suc n} = (f (Suc n) \<otimes> finprod G f {..n})" |
477 |
by (simp add: Pi_def atMost_Suc) |
|
478 |
||
27933 | 479 |
lemma finprod_Suc2: |
14750 | 480 |
"f \<in> {..Suc n} -> carrier G ==> |
13936 | 481 |
finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} \<otimes> f 0)" |
482 |
proof (induct n) |
|
483 |
case 0 thus ?case by (simp add: Pi_def) |
|
484 |
next |
|
485 |
case Suc thus ?case by (simp add: m_assoc Pi_def) |
|
486 |
qed |
|
487 |
||
27933 | 488 |
lemma finprod_mult [simp]: |
14750 | 489 |
"[| f \<in> {..n} -> carrier G; g \<in> {..n} -> carrier G |] ==> |
13936 | 490 |
finprod G (%i. f i \<otimes> g i) {..n::nat} = |
491 |
finprod G f {..n} \<otimes> finprod G g {..n}" |
|
492 |
by (induct n) (simp_all add: m_ac Pi_def) |
|
493 |
||
27699 | 494 |
(* The following two were contributed by Jeremy Avigad. *) |
495 |
||
27933 | 496 |
lemma finprod_reindex: |
27699 | 497 |
assumes fin: "finite A" |
498 |
shows "f : (h ` A) \<rightarrow> carrier G \<Longrightarrow> |
|
499 |
inj_on h A ==> finprod G f (h ` A) = finprod G (%x. f (h x)) A" |
|
46721 | 500 |
using fin |
501 |
by induct (auto simp add: Pi_def) |
|
27699 | 502 |
|
27933 | 503 |
lemma finprod_const: |
27699 | 504 |
assumes fin [simp]: "finite A" |
505 |
and a [simp]: "a : carrier G" |
|
506 |
shows "finprod G (%x. a) A = a (^) card A" |
|
507 |
using fin apply induct |
|
508 |
apply force |
|
509 |
apply (subst finprod_insert) |
|
510 |
apply auto |
|
511 |
apply (subst m_comm) |
|
512 |
apply auto |
|
46721 | 513 |
done |
27699 | 514 |
|
27933 | 515 |
(* The following lemma was contributed by Jesus Aransay. *) |
516 |
||
517 |
lemma finprod_singleton: |
|
518 |
assumes i_in_A: "i \<in> A" and fin_A: "finite A" and f_Pi: "f \<in> A \<rightarrow> carrier G" |
|
519 |
shows "(\<Otimes>j\<in>A. if i = j then f j else \<one>) = f i" |
|
29237 | 520 |
using i_in_A finprod_insert [of "A - {i}" i "(\<lambda>j. if i = j then f j else \<one>)"] |
521 |
fin_A f_Pi finprod_one [of "A - {i}"] |
|
522 |
finprod_cong [of "A - {i}" "A - {i}" "(\<lambda>j. if i = j then f j else \<one>)" "(\<lambda>i. \<one>)"] |
|
27933 | 523 |
unfolding Pi_def simp_implies_def by (force simp add: insert_absorb) |
524 |
||
13936 | 525 |
end |
27933 | 526 |
|
527 |
end |