author | wenzelm |
Tue, 07 Mar 2017 17:21:41 +0100 | |
changeset 65143 | 36cd85caf09a |
parent 63167 | 0909deb8059b |
child 67091 | 1393c2340eec |
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 |
14750 | 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 |
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 |
58860 | 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 |
|
61382 | 64 |
text \<open>Left-Commutative Operations\<close> |
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]: |
|
58860 | 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 |
|
63167 | 114 |
txt \<open>force simplification of \<open>card A < card (insert ...)\<close>.\<close> |
13936 | 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 |
|
61382 | 122 |
txt \<open>case @{prop "xa \<notin> xb"}.\<close> |
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 |
||
63167 | 224 |
\<comment> \<open>Delete rules to do with \<open>foldSetD\<close> relation.\<close> |
13936 | 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 |
|
61382 | 233 |
text \<open>Commutative Monoids\<close> |
13936 | 234 |
|
61382 | 235 |
text \<open> |
63167 | 236 |
We enter a more restrictive context, with \<open>f :: 'a => 'a => 'a\<close> |
237 |
instead of \<open>'b => 'a => 'a\<close>. |
|
61382 | 238 |
\<close> |
13936 | 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) |
|
61382 | 266 |
with \<open>x \<in> D\<close> 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 |
|
61382 | 288 |
subsubsection \<open>Products over Finite Sets\<close> |
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 |
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
|
295 |
else \<one>\<^bsub>G\<^esub>)" |
13936 | 296 |
|
14651 | 297 |
syntax |
298 |
"_finprod" :: "index => idt => 'a set => 'b => 'b" |
|
14666 | 299 |
("(3\<Otimes>__\<in>_. _)" [1000, 0, 51, 10] 10) |
14651 | 300 |
translations |
62105 | 301 |
"\<Otimes>\<^bsub>G\<^esub>i\<in>A. b" \<rightleftharpoons> "CONST finprod G (%i. b) A" |
63167 | 302 |
\<comment> \<open>Beware of argument permutation!\<close> |
13936 | 303 |
|
304 |
lemma (in comm_monoid) finprod_empty [simp]: |
|
305 |
"finprod G f {} = \<one>" |
|
306 |
by (simp add: finprod_def) |
|
307 |
||
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
|
308 |
lemma (in comm_monoid) finprod_infinite[simp]: |
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
|
309 |
"\<not> finite A \<Longrightarrow> finprod G f A = \<one>" |
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
|
310 |
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
|
311 |
|
13936 | 312 |
declare funcsetI [intro] |
313 |
funcset_mem [dest] |
|
314 |
||
27933 | 315 |
context comm_monoid begin |
316 |
||
317 |
lemma finprod_insert [simp]: |
|
61384 | 318 |
"[| finite F; a \<notin> F; f \<in> F \<rightarrow> carrier G; f a \<in> carrier G |] ==> |
13936 | 319 |
finprod G f (insert a F) = f a \<otimes> finprod G f F" |
320 |
apply (rule trans) |
|
321 |
apply (simp add: finprod_def) |
|
322 |
apply (rule trans) |
|
323 |
apply (rule LCD.foldD_insert [OF LCD.intro [of "insert a F"]]) |
|
324 |
apply simp |
|
325 |
apply (rule m_lcomm) |
|
326 |
apply fast |
|
327 |
apply fast |
|
328 |
apply assumption |
|
46721 | 329 |
apply fastforce |
13936 | 330 |
apply simp+ |
331 |
apply fast |
|
332 |
apply (auto simp add: finprod_def) |
|
333 |
done |
|
334 |
||
60773 | 335 |
lemma finprod_one [simp]: "(\<Otimes>i\<in>A. \<one>) = \<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
|
336 |
proof (induct A rule: infinite_finite_induct) |
13936 | 337 |
case empty show ?case by simp |
338 |
next |
|
15328 | 339 |
case (insert a A) |
61384 | 340 |
have "(%i. \<one>) \<in> A \<rightarrow> carrier G" by auto |
13936 | 341 |
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
|
342 |
qed simp |
13936 | 343 |
|
27933 | 344 |
lemma finprod_closed [simp]: |
13936 | 345 |
fixes A |
61384 | 346 |
assumes f: "f \<in> A \<rightarrow> carrier G" |
13936 | 347 |
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
|
348 |
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
|
349 |
proof (induct A rule: infinite_finite_induct) |
13936 | 350 |
case empty show ?case by simp |
351 |
next |
|
15328 | 352 |
case (insert a A) |
13936 | 353 |
then have a: "f a \<in> carrier G" by fast |
61384 | 354 |
from insert have A: "f \<in> A \<rightarrow> carrier G" by fast |
13936 | 355 |
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
|
356 |
qed simp |
13936 | 357 |
|
358 |
lemma funcset_Int_left [simp, intro]: |
|
61384 | 359 |
"[| f \<in> A \<rightarrow> C; f \<in> B \<rightarrow> C |] ==> f \<in> A Int B \<rightarrow> C" |
13936 | 360 |
by fast |
361 |
||
362 |
lemma funcset_Un_left [iff]: |
|
61384 | 363 |
"(f \<in> A Un B \<rightarrow> C) = (f \<in> A \<rightarrow> C & f \<in> B \<rightarrow> C)" |
13936 | 364 |
by fast |
365 |
||
27933 | 366 |
lemma finprod_Un_Int: |
61384 | 367 |
"[| finite A; finite B; g \<in> A \<rightarrow> carrier G; g \<in> B \<rightarrow> carrier G |] ==> |
13936 | 368 |
finprod G g (A Un B) \<otimes> finprod G g (A Int B) = |
369 |
finprod G g A \<otimes> finprod G g B" |
|
63167 | 370 |
\<comment> \<open>The reversed orientation looks more natural, but LOOPS as a simprule!\<close> |
22265 | 371 |
proof (induct set: finite) |
46721 | 372 |
case empty then show ?case by simp |
13936 | 373 |
next |
15328 | 374 |
case (insert a A) |
13936 | 375 |
then have a: "g a \<in> carrier G" by fast |
61384 | 376 |
from insert have A: "g \<in> A \<rightarrow> carrier G" by fast |
13936 | 377 |
from insert A a show ?case |
46721 | 378 |
by (simp add: m_ac Int_insert_left insert_absorb Int_mono2) |
13936 | 379 |
qed |
380 |
||
27933 | 381 |
lemma finprod_Un_disjoint: |
13936 | 382 |
"[| finite A; finite B; A Int B = {}; |
61384 | 383 |
g \<in> A \<rightarrow> carrier G; g \<in> B \<rightarrow> carrier G |] |
13936 | 384 |
==> finprod G g (A Un B) = finprod G g A \<otimes> finprod G g B" |
385 |
apply (subst finprod_Un_Int [symmetric]) |
|
46721 | 386 |
apply auto |
13936 | 387 |
done |
388 |
||
27933 | 389 |
lemma finprod_multf: |
61384 | 390 |
"[| f \<in> A \<rightarrow> carrier G; g \<in> A \<rightarrow> carrier G |] ==> |
13936 | 391 |
finprod G (%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
|
392 |
proof (induct A rule: infinite_finite_induct) |
13936 | 393 |
case empty show ?case by simp |
394 |
next |
|
15328 | 395 |
case (insert a A) then |
61384 | 396 |
have fA: "f \<in> A \<rightarrow> carrier G" by fast |
14750 | 397 |
from insert have fa: "f a \<in> carrier G" by fast |
61384 | 398 |
from insert have gA: "g \<in> A \<rightarrow> carrier G" by fast |
14750 | 399 |
from insert have ga: "g a \<in> carrier G" by fast |
61384 | 400 |
from insert have fgA: "(%x. f x \<otimes> g x) \<in> A \<rightarrow> carrier G" |
13936 | 401 |
by (simp add: Pi_def) |
15095
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
402 |
show ?case |
63f5f4c265dd
Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents:
14750
diff
changeset
|
403 |
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
|
404 |
qed simp |
13936 | 405 |
|
27933 | 406 |
lemma finprod_cong': |
61384 | 407 |
"[| A = B; g \<in> B \<rightarrow> carrier G; |
14750 | 408 |
!!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B" |
13936 | 409 |
proof - |
61384 | 410 |
assume prems: "A = B" "g \<in> B \<rightarrow> carrier G" |
14750 | 411 |
"!!i. i \<in> B ==> f i = g i" |
13936 | 412 |
show ?thesis |
413 |
proof (cases "finite B") |
|
414 |
case True |
|
61384 | 415 |
then have "!!A. [| A = B; g \<in> B \<rightarrow> carrier G; |
14750 | 416 |
!!i. i \<in> B ==> f i = g i |] ==> finprod G f A = finprod G g B" |
13936 | 417 |
proof induct |
418 |
case empty thus ?case by simp |
|
419 |
next |
|
15328 | 420 |
case (insert x B) |
13936 | 421 |
then have "finprod G f A = finprod G f (insert x B)" by simp |
422 |
also from insert have "... = f x \<otimes> finprod G f B" |
|
423 |
proof (intro finprod_insert) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
424 |
show "finite B" by fact |
13936 | 425 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
426 |
show "x ~: B" by fact |
13936 | 427 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
428 |
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
|
429 |
"g \<in> insert x B \<rightarrow> carrier G" |
61384 | 430 |
thus "f \<in> B \<rightarrow> carrier G" by fastforce |
13936 | 431 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
32693
diff
changeset
|
432 |
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
|
433 |
"g \<in> insert x B \<rightarrow> carrier G" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
434 |
thus "f x \<in> carrier G" by fastforce |
13936 | 435 |
qed |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
41433
diff
changeset
|
436 |
also from insert have "... = g x \<otimes> finprod G g B" by fastforce |
13936 | 437 |
also from insert have "... = finprod G g (insert x B)" |
438 |
by (intro finprod_insert [THEN sym]) auto |
|
439 |
finally show ?case . |
|
440 |
qed |
|
441 |
with prems show ?thesis by simp |
|
442 |
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
|
443 |
case False with prems show ?thesis by simp |
13936 | 444 |
qed |
445 |
qed |
|
446 |
||
27933 | 447 |
lemma finprod_cong: |
61384 | 448 |
"[| A = B; f \<in> B \<rightarrow> carrier G = True; |
41433
1b8ff770f02c
Abelian group facts obtained from group facts via interpretation (sublocale).
ballarin
parents:
40786
diff
changeset
|
449 |
!!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
|
450 |
(* 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
|
451 |
by (rule finprod_cong') (auto simp add: simp_implies_def) |
13936 | 452 |
|
61382 | 453 |
text \<open>Usually, if this rule causes a failed congruence proof error, |
63167 | 454 |
the reason is that the premise \<open>g \<in> B \<rightarrow> carrier G\<close> cannot be shown. |
13936 | 455 |
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
|
456 |
For this reason, @{thm [source] finprod_cong} |
13936 | 457 |
is not added to the simpset by default. |
61382 | 458 |
\<close> |
13936 | 459 |
|
27933 | 460 |
end |
461 |
||
13936 | 462 |
declare funcsetI [rule del] |
463 |
funcset_mem [rule del] |
|
464 |
||
27933 | 465 |
context comm_monoid begin |
466 |
||
467 |
lemma finprod_0 [simp]: |
|
61384 | 468 |
"f \<in> {0::nat} \<rightarrow> carrier G ==> finprod G f {..0} = f 0" |
13936 | 469 |
by (simp add: Pi_def) |
470 |
||
27933 | 471 |
lemma finprod_Suc [simp]: |
61384 | 472 |
"f \<in> {..Suc n} \<rightarrow> carrier G ==> |
13936 | 473 |
finprod G f {..Suc n} = (f (Suc n) \<otimes> finprod G f {..n})" |
474 |
by (simp add: Pi_def atMost_Suc) |
|
475 |
||
27933 | 476 |
lemma finprod_Suc2: |
61384 | 477 |
"f \<in> {..Suc n} \<rightarrow> carrier G ==> |
13936 | 478 |
finprod G f {..Suc n} = (finprod G (%i. f (Suc i)) {..n} \<otimes> f 0)" |
479 |
proof (induct n) |
|
480 |
case 0 thus ?case by (simp add: Pi_def) |
|
481 |
next |
|
482 |
case Suc thus ?case by (simp add: m_assoc Pi_def) |
|
483 |
qed |
|
484 |
||
27933 | 485 |
lemma finprod_mult [simp]: |
61384 | 486 |
"[| f \<in> {..n} \<rightarrow> carrier G; g \<in> {..n} \<rightarrow> carrier G |] ==> |
13936 | 487 |
finprod G (%i. f i \<otimes> g i) {..n::nat} = |
488 |
finprod G f {..n} \<otimes> finprod G g {..n}" |
|
489 |
by (induct n) (simp_all add: m_ac Pi_def) |
|
490 |
||
27699 | 491 |
(* The following two were contributed by Jeremy Avigad. *) |
492 |
||
27933 | 493 |
lemma finprod_reindex: |
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
|
494 |
"f : (h ` A) \<rightarrow> carrier G \<Longrightarrow> |
27699 | 495 |
inj_on h A ==> finprod G f (h ` A) = finprod G (%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
|
496 |
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
|
497 |
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
|
498 |
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
|
499 |
using finite_imageD by blast |
61382 | 500 |
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
|
501 |
qed (auto simp add: Pi_def) |
27699 | 502 |
|
27933 | 503 |
lemma finprod_const: |
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
|
504 |
assumes a [simp]: "a : carrier G" |
27699 | 505 |
shows "finprod G (%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
|
506 |
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
|
507 |
case (insert b 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
|
508 |
show ?case |
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
|
509 |
proof (subst finprod_insert[OF insert(1-2)]) |
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
|
510 |
show "a \<otimes> (\<Otimes>x\<in>A. a) = a (^) card (insert b 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
|
511 |
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
|
512 |
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
|
513 |
qed auto |
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 |