| author | paulson |
| Mon, 15 Mar 2004 10:58:49 +0100 | |
| changeset 14470 | 1ffe42cfaefe |
| parent 14449 | d5c3d21df790 |
| child 14485 | ea2707645af8 |
| permissions | -rw-r--r-- |
| 12396 | 1 |
(* Title: HOL/Finite_Set.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
4 |
Additions by Jeremy Avigad in Feb 2004 |
| 12396 | 5 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
6 |
*) |
|
7 |
||
8 |
header {* Finite sets *}
|
|
9 |
||
10 |
theory Finite_Set = Divides + Power + Inductive + SetInterval: |
|
11 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
12 |
subsection {* Types Classes @{text plus_ac0} and @{text times_ac1} *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
13 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
14 |
axclass plus_ac0 < plus, zero |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
15 |
commute: "x + y = y + x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
16 |
assoc: "(x + y) + z = x + (y + z)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
17 |
zero [simp]: "0 + x = x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
18 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
19 |
lemma plus_ac0_left_commute: "x + (y+z) = y + ((x+z)::'a::plus_ac0)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
20 |
apply (rule plus_ac0.commute [THEN trans]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
21 |
apply (rule plus_ac0.assoc [THEN trans]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
22 |
apply (rule plus_ac0.commute [THEN arg_cong]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
23 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
24 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
25 |
lemma plus_ac0_zero_right [simp]: "x + 0 = (x ::'a::plus_ac0)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
26 |
apply (rule plus_ac0.commute [THEN trans]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
27 |
apply (rule plus_ac0.zero) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
28 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
29 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
30 |
lemmas plus_ac0 = plus_ac0.assoc plus_ac0.commute plus_ac0_left_commute |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
31 |
plus_ac0.zero plus_ac0_zero_right |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
32 |
|
| 14449 | 33 |
instance almost_semiring < plus_ac0 |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
34 |
proof qed (rule add_commute add_assoc almost_semiring.add_0)+ |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
35 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
36 |
axclass times_ac1 < times, one |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
37 |
commute: "x * y = y * x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
38 |
assoc: "(x * y) * z = x * (y * z)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
39 |
one: "1 * x = x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
40 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
41 |
theorem times_ac1_left_commute: "(x::'a::times_ac1) * ((y::'a) * z) = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
42 |
y * (x * z)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
43 |
proof - |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
44 |
have "(x::'a::times_ac1) * (y * z) = (x * y) * z" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
45 |
by (rule times_ac1.assoc [THEN sym]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
46 |
also have "x * y = y * x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
47 |
by (rule times_ac1.commute) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
48 |
also have "(y * x) * z = y * (x * z)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
49 |
by (rule times_ac1.assoc) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
50 |
finally show ?thesis . |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
51 |
qed |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
52 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
53 |
theorem times_ac1_one_right: "(x::'a::times_ac1) * 1 = x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
54 |
proof - |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
55 |
have "x * 1 = 1 * x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
56 |
by (rule times_ac1.commute) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
57 |
also have "... = x" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
58 |
by (rule times_ac1.one) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
59 |
finally show ?thesis . |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
60 |
qed |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
61 |
|
| 14449 | 62 |
instance almost_semiring < times_ac1 |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
63 |
apply intro_classes |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
64 |
apply (rule mult_commute) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
65 |
apply (rule mult_assoc, simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
66 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
67 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
68 |
theorems times_ac1 = times_ac1.assoc times_ac1.commute times_ac1_left_commute |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
69 |
times_ac1.one times_ac1_one_right |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
70 |
|
| 12396 | 71 |
subsection {* Collection of finite sets *}
|
72 |
||
73 |
consts Finites :: "'a set set" |
|
| 13737 | 74 |
syntax |
75 |
finite :: "'a set => bool" |
|
76 |
translations |
|
77 |
"finite A" == "A : Finites" |
|
| 12396 | 78 |
|
79 |
inductive Finites |
|
80 |
intros |
|
81 |
emptyI [simp, intro!]: "{} : Finites"
|
|
82 |
insertI [simp, intro!]: "A : Finites ==> insert a A : Finites" |
|
83 |
||
84 |
axclass finite \<subseteq> type |
|
85 |
finite: "finite UNIV" |
|
86 |
||
| 13737 | 87 |
lemma ex_new_if_finite: -- "does not depend on def of finite at all" |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
88 |
"[| ~finite(UNIV::'a set); finite A |] ==> \<exists>a::'a. a \<notin> A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
89 |
by(subgoal_tac "A \<noteq> UNIV", blast, blast) |
| 13737 | 90 |
|
| 12396 | 91 |
|
92 |
lemma finite_induct [case_names empty insert, induct set: Finites]: |
|
93 |
"finite F ==> |
|
94 |
P {} ==> (!!F x. finite F ==> x \<notin> F ==> P F ==> P (insert x F)) ==> P F"
|
|
95 |
-- {* Discharging @{text "x \<notin> F"} entails extra work. *}
|
|
96 |
proof - |
|
| 13421 | 97 |
assume "P {}" and
|
98 |
insert: "!!F x. finite F ==> x \<notin> F ==> P F ==> P (insert x F)" |
|
| 12396 | 99 |
assume "finite F" |
100 |
thus "P F" |
|
101 |
proof induct |
|
102 |
show "P {}" .
|
|
103 |
fix F x assume F: "finite F" and P: "P F" |
|
104 |
show "P (insert x F)" |
|
105 |
proof cases |
|
106 |
assume "x \<in> F" |
|
107 |
hence "insert x F = F" by (rule insert_absorb) |
|
108 |
with P show ?thesis by (simp only:) |
|
109 |
next |
|
110 |
assume "x \<notin> F" |
|
111 |
from F this P show ?thesis by (rule insert) |
|
112 |
qed |
|
113 |
qed |
|
114 |
qed |
|
115 |
||
116 |
lemma finite_subset_induct [consumes 2, case_names empty insert]: |
|
117 |
"finite F ==> F \<subseteq> A ==> |
|
118 |
P {} ==> (!!F a. finite F ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)) ==>
|
|
119 |
P F" |
|
120 |
proof - |
|
| 13421 | 121 |
assume "P {}" and insert:
|
122 |
"!!F a. finite F ==> a \<in> A ==> a \<notin> F ==> P F ==> P (insert a F)" |
|
| 12396 | 123 |
assume "finite F" |
124 |
thus "F \<subseteq> A ==> P F" |
|
125 |
proof induct |
|
126 |
show "P {}" .
|
|
127 |
fix F x assume "finite F" and "x \<notin> F" |
|
128 |
and P: "F \<subseteq> A ==> P F" and i: "insert x F \<subseteq> A" |
|
129 |
show "P (insert x F)" |
|
130 |
proof (rule insert) |
|
131 |
from i show "x \<in> A" by blast |
|
132 |
from i have "F \<subseteq> A" by blast |
|
133 |
with P show "P F" . |
|
134 |
qed |
|
135 |
qed |
|
136 |
qed |
|
137 |
||
138 |
lemma finite_UnI: "finite F ==> finite G ==> finite (F Un G)" |
|
139 |
-- {* The union of two finite sets is finite. *}
|
|
140 |
by (induct set: Finites) simp_all |
|
141 |
||
142 |
lemma finite_subset: "A \<subseteq> B ==> finite B ==> finite A" |
|
143 |
-- {* Every subset of a finite set is finite. *}
|
|
144 |
proof - |
|
145 |
assume "finite B" |
|
146 |
thus "!!A. A \<subseteq> B ==> finite A" |
|
147 |
proof induct |
|
148 |
case empty |
|
149 |
thus ?case by simp |
|
150 |
next |
|
151 |
case (insert F x A) |
|
152 |
have A: "A \<subseteq> insert x F" and r: "A - {x} \<subseteq> F ==> finite (A - {x})" .
|
|
153 |
show "finite A" |
|
154 |
proof cases |
|
155 |
assume x: "x \<in> A" |
|
156 |
with A have "A - {x} \<subseteq> F" by (simp add: subset_insert_iff)
|
|
157 |
with r have "finite (A - {x})" .
|
|
158 |
hence "finite (insert x (A - {x}))" ..
|
|
159 |
also have "insert x (A - {x}) = A" by (rule insert_Diff)
|
|
160 |
finally show ?thesis . |
|
161 |
next |
|
162 |
show "A \<subseteq> F ==> ?thesis" . |
|
163 |
assume "x \<notin> A" |
|
164 |
with A show "A \<subseteq> F" by (simp add: subset_insert_iff) |
|
165 |
qed |
|
166 |
qed |
|
167 |
qed |
|
168 |
||
169 |
lemma finite_Un [iff]: "finite (F Un G) = (finite F & finite G)" |
|
170 |
by (blast intro: finite_subset [of _ "X Un Y", standard] finite_UnI) |
|
171 |
||
172 |
lemma finite_Int [simp, intro]: "finite F | finite G ==> finite (F Int G)" |
|
173 |
-- {* The converse obviously fails. *}
|
|
174 |
by (blast intro: finite_subset) |
|
175 |
||
176 |
lemma finite_insert [simp]: "finite (insert a A) = finite A" |
|
177 |
apply (subst insert_is_Un) |
|
| 14208 | 178 |
apply (simp only: finite_Un, blast) |
| 12396 | 179 |
done |
180 |
||
181 |
lemma finite_empty_induct: |
|
182 |
"finite A ==> |
|
183 |
P A ==> (!!a A. finite A ==> a:A ==> P A ==> P (A - {a})) ==> P {}"
|
|
184 |
proof - |
|
185 |
assume "finite A" |
|
186 |
and "P A" and "!!a A. finite A ==> a:A ==> P A ==> P (A - {a})"
|
|
187 |
have "P (A - A)" |
|
188 |
proof - |
|
189 |
fix c b :: "'a set" |
|
190 |
presume c: "finite c" and b: "finite b" |
|
191 |
and P1: "P b" and P2: "!!x y. finite y ==> x \<in> y ==> P y ==> P (y - {x})"
|
|
192 |
from c show "c \<subseteq> b ==> P (b - c)" |
|
193 |
proof induct |
|
194 |
case empty |
|
195 |
from P1 show ?case by simp |
|
196 |
next |
|
197 |
case (insert F x) |
|
198 |
have "P (b - F - {x})"
|
|
199 |
proof (rule P2) |
|
200 |
from _ b show "finite (b - F)" by (rule finite_subset) blast |
|
201 |
from insert show "x \<in> b - F" by simp |
|
202 |
from insert show "P (b - F)" by simp |
|
203 |
qed |
|
204 |
also have "b - F - {x} = b - insert x F" by (rule Diff_insert [symmetric])
|
|
205 |
finally show ?case . |
|
206 |
qed |
|
207 |
next |
|
208 |
show "A \<subseteq> A" .. |
|
209 |
qed |
|
210 |
thus "P {}" by simp
|
|
211 |
qed |
|
212 |
||
213 |
lemma finite_Diff [simp]: "finite B ==> finite (B - Ba)" |
|
214 |
by (rule Diff_subset [THEN finite_subset]) |
|
215 |
||
216 |
lemma finite_Diff_insert [iff]: "finite (A - insert a B) = finite (A - B)" |
|
217 |
apply (subst Diff_insert) |
|
218 |
apply (case_tac "a : A - B") |
|
219 |
apply (rule finite_insert [symmetric, THEN trans]) |
|
| 14208 | 220 |
apply (subst insert_Diff, simp_all) |
| 12396 | 221 |
done |
222 |
||
223 |
||
| 13825 | 224 |
subsubsection {* Image and Inverse Image over Finite Sets *}
|
225 |
||
226 |
lemma finite_imageI[simp]: "finite F ==> finite (h ` F)" |
|
227 |
-- {* The image of a finite set is finite. *}
|
|
228 |
by (induct set: Finites) simp_all |
|
229 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
230 |
lemma finite_surj: "finite A ==> B <= f ` A ==> finite B" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
231 |
apply (frule finite_imageI) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
232 |
apply (erule finite_subset, assumption) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
233 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
234 |
|
| 13825 | 235 |
lemma finite_range_imageI: |
236 |
"finite (range g) ==> finite (range (%x. f (g x)))" |
|
| 14208 | 237 |
apply (drule finite_imageI, simp) |
| 13825 | 238 |
done |
239 |
||
| 12396 | 240 |
lemma finite_imageD: "finite (f`A) ==> inj_on f A ==> finite A" |
241 |
proof - |
|
242 |
have aux: "!!A. finite (A - {}) = finite A" by simp
|
|
243 |
fix B :: "'a set" |
|
244 |
assume "finite B" |
|
245 |
thus "!!A. f`A = B ==> inj_on f A ==> finite A" |
|
246 |
apply induct |
|
247 |
apply simp |
|
248 |
apply (subgoal_tac "EX y:A. f y = x & F = f ` (A - {y})")
|
|
249 |
apply clarify |
|
250 |
apply (simp (no_asm_use) add: inj_on_def) |
|
| 14208 | 251 |
apply (blast dest!: aux [THEN iffD1], atomize) |
| 12396 | 252 |
apply (erule_tac V = "ALL A. ?PP (A)" in thin_rl) |
| 14208 | 253 |
apply (frule subsetD [OF equalityD2 insertI1], clarify) |
| 12396 | 254 |
apply (rule_tac x = xa in bexI) |
255 |
apply (simp_all add: inj_on_image_set_diff) |
|
256 |
done |
|
257 |
qed (rule refl) |
|
258 |
||
259 |
||
| 13825 | 260 |
lemma inj_vimage_singleton: "inj f ==> f-`{a} \<subseteq> {THE x. f x = a}"
|
261 |
-- {* The inverse image of a singleton under an injective function
|
|
262 |
is included in a singleton. *} |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
263 |
apply (auto simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
264 |
apply (blast intro: the_equality [symmetric]) |
| 13825 | 265 |
done |
266 |
||
267 |
lemma finite_vimageI: "[|finite F; inj h|] ==> finite (h -` F)" |
|
268 |
-- {* The inverse image of a finite set under an injective function
|
|
269 |
is finite. *} |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
270 |
apply (induct set: Finites, simp_all) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
271 |
apply (subst vimage_insert) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
272 |
apply (simp add: finite_Un finite_subset [OF inj_vimage_singleton]) |
| 13825 | 273 |
done |
274 |
||
275 |
||
| 12396 | 276 |
subsubsection {* The finite UNION of finite sets *}
|
277 |
||
278 |
lemma finite_UN_I: "finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (UN a:A. B a)" |
|
279 |
by (induct set: Finites) simp_all |
|
280 |
||
281 |
text {*
|
|
282 |
Strengthen RHS to |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
283 |
@{prop "((ALL x:A. finite (B x)) & finite {x. x:A & B x \<noteq> {}})"}?
|
| 12396 | 284 |
|
285 |
We'd need to prove |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
286 |
@{prop "finite C ==> ALL A B. (UNION A B) <= C --> finite {x. x:A & B x \<noteq> {}}"}
|
| 12396 | 287 |
by induction. *} |
288 |
||
289 |
lemma finite_UN [simp]: "finite A ==> finite (UNION A B) = (ALL x:A. finite (B x))" |
|
290 |
by (blast intro: finite_UN_I finite_subset) |
|
291 |
||
292 |
||
293 |
subsubsection {* Sigma of finite sets *}
|
|
294 |
||
295 |
lemma finite_SigmaI [simp]: |
|
296 |
"finite A ==> (!!a. a:A ==> finite (B a)) ==> finite (SIGMA a:A. B a)" |
|
297 |
by (unfold Sigma_def) (blast intro!: finite_UN_I) |
|
298 |
||
299 |
lemma finite_Prod_UNIV: |
|
300 |
"finite (UNIV::'a set) ==> finite (UNIV::'b set) ==> finite (UNIV::('a * 'b) set)"
|
|
301 |
apply (subgoal_tac "(UNIV:: ('a * 'b) set) = Sigma UNIV (%x. UNIV)")
|
|
302 |
apply (erule ssubst) |
|
| 14208 | 303 |
apply (erule finite_SigmaI, auto) |
| 12396 | 304 |
done |
305 |
||
306 |
instance unit :: finite |
|
307 |
proof |
|
308 |
have "finite {()}" by simp
|
|
309 |
also have "{()} = UNIV" by auto
|
|
310 |
finally show "finite (UNIV :: unit set)" . |
|
311 |
qed |
|
312 |
||
313 |
instance * :: (finite, finite) finite |
|
314 |
proof |
|
315 |
show "finite (UNIV :: ('a \<times> 'b) set)"
|
|
316 |
proof (rule finite_Prod_UNIV) |
|
317 |
show "finite (UNIV :: 'a set)" by (rule finite) |
|
318 |
show "finite (UNIV :: 'b set)" by (rule finite) |
|
319 |
qed |
|
320 |
qed |
|
321 |
||
322 |
||
323 |
subsubsection {* The powerset of a finite set *}
|
|
324 |
||
325 |
lemma finite_Pow_iff [iff]: "finite (Pow A) = finite A" |
|
326 |
proof |
|
327 |
assume "finite (Pow A)" |
|
328 |
with _ have "finite ((%x. {x}) ` A)" by (rule finite_subset) blast
|
|
329 |
thus "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp |
|
330 |
next |
|
331 |
assume "finite A" |
|
332 |
thus "finite (Pow A)" |
|
333 |
by induct (simp_all add: finite_UnI finite_imageI Pow_insert) |
|
334 |
qed |
|
335 |
||
336 |
lemma finite_converse [iff]: "finite (r^-1) = finite r" |
|
337 |
apply (subgoal_tac "r^-1 = (%(x,y). (y,x))`r") |
|
338 |
apply simp |
|
339 |
apply (rule iffI) |
|
340 |
apply (erule finite_imageD [unfolded inj_on_def]) |
|
341 |
apply (simp split add: split_split) |
|
342 |
apply (erule finite_imageI) |
|
| 14208 | 343 |
apply (simp add: converse_def image_def, auto) |
| 12396 | 344 |
apply (rule bexI) |
345 |
prefer 2 apply assumption |
|
346 |
apply simp |
|
347 |
done |
|
348 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
349 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
350 |
subsubsection {* Intervals of nats *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
351 |
|
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
352 |
lemma finite_lessThan [iff]: fixes k :: nat shows "finite {..k(}"
|
| 12396 | 353 |
by (induct k) (simp_all add: lessThan_Suc) |
354 |
||
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
355 |
lemma finite_atMost [iff]: fixes k :: nat shows "finite {..k}"
|
| 12396 | 356 |
by (induct k) (simp_all add: atMost_Suc) |
357 |
||
| 13735 | 358 |
lemma finite_greaterThanLessThan [iff]: |
359 |
fixes l :: nat shows "finite {)l..u(}"
|
|
360 |
by (simp add: greaterThanLessThan_def) |
|
361 |
||
362 |
lemma finite_atLeastLessThan [iff]: |
|
363 |
fixes l :: nat shows "finite {l..u(}"
|
|
364 |
by (simp add: atLeastLessThan_def) |
|
365 |
||
366 |
lemma finite_greaterThanAtMost [iff]: |
|
367 |
fixes l :: nat shows "finite {)l..u}"
|
|
368 |
by (simp add: greaterThanAtMost_def) |
|
369 |
||
370 |
lemma finite_atLeastAtMost [iff]: |
|
371 |
fixes l :: nat shows "finite {l..u}"
|
|
372 |
by (simp add: atLeastAtMost_def) |
|
373 |
||
| 12396 | 374 |
lemma bounded_nat_set_is_finite: |
375 |
"(ALL i:N. i < (n::nat)) ==> finite N" |
|
376 |
-- {* A bounded set of natural numbers is finite. *}
|
|
377 |
apply (rule finite_subset) |
|
| 14208 | 378 |
apply (rule_tac [2] finite_lessThan, auto) |
| 12396 | 379 |
done |
380 |
||
381 |
||
382 |
subsubsection {* Finiteness of transitive closure *}
|
|
383 |
||
384 |
text {* (Thanks to Sidi Ehmety) *}
|
|
385 |
||
386 |
lemma finite_Field: "finite r ==> finite (Field r)" |
|
387 |
-- {* A finite relation has a finite field (@{text "= domain \<union> range"}. *}
|
|
388 |
apply (induct set: Finites) |
|
389 |
apply (auto simp add: Field_def Domain_insert Range_insert) |
|
390 |
done |
|
391 |
||
392 |
lemma trancl_subset_Field2: "r^+ <= Field r \<times> Field r" |
|
393 |
apply clarify |
|
394 |
apply (erule trancl_induct) |
|
395 |
apply (auto simp add: Field_def) |
|
396 |
done |
|
397 |
||
398 |
lemma finite_trancl: "finite (r^+) = finite r" |
|
399 |
apply auto |
|
400 |
prefer 2 |
|
401 |
apply (rule trancl_subset_Field2 [THEN finite_subset]) |
|
402 |
apply (rule finite_SigmaI) |
|
403 |
prefer 3 |
|
|
13704
854501b1e957
Transitive closure is now defined inductively as well.
berghofe
parents:
13595
diff
changeset
|
404 |
apply (blast intro: r_into_trancl' finite_subset) |
| 12396 | 405 |
apply (auto simp add: finite_Field) |
406 |
done |
|
407 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
408 |
lemma finite_cartesian_product: "[| finite A; finite B |] ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
409 |
finite (A <*> B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
410 |
by (rule finite_SigmaI) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
411 |
|
| 12396 | 412 |
|
413 |
subsection {* Finite cardinality *}
|
|
414 |
||
415 |
text {*
|
|
416 |
This definition, although traditional, is ugly to work with: @{text
|
|
417 |
"card A == LEAST n. EX f. A = {f i | i. i < n}"}. Therefore we have
|
|
418 |
switched to an inductive one: |
|
419 |
*} |
|
420 |
||
421 |
consts cardR :: "('a set \<times> nat) set"
|
|
422 |
||
423 |
inductive cardR |
|
424 |
intros |
|
425 |
EmptyI: "({}, 0) : cardR"
|
|
426 |
InsertI: "(A, n) : cardR ==> a \<notin> A ==> (insert a A, Suc n) : cardR" |
|
427 |
||
428 |
constdefs |
|
429 |
card :: "'a set => nat" |
|
430 |
"card A == THE n. (A, n) : cardR" |
|
431 |
||
432 |
inductive_cases cardR_emptyE: "({}, n) : cardR"
|
|
433 |
inductive_cases cardR_insertE: "(insert a A,n) : cardR" |
|
434 |
||
435 |
lemma cardR_SucD: "(A, n) : cardR ==> a : A --> (EX m. n = Suc m)" |
|
436 |
by (induct set: cardR) simp_all |
|
437 |
||
438 |
lemma cardR_determ_aux1: |
|
439 |
"(A, m): cardR ==> (!!n a. m = Suc n ==> a:A ==> (A - {a}, n) : cardR)"
|
|
| 14208 | 440 |
apply (induct set: cardR, auto) |
441 |
apply (simp add: insert_Diff_if, auto) |
|
| 12396 | 442 |
apply (drule cardR_SucD) |
443 |
apply (blast intro!: cardR.intros) |
|
444 |
done |
|
445 |
||
446 |
lemma cardR_determ_aux2: "(insert a A, Suc m) : cardR ==> a \<notin> A ==> (A, m) : cardR" |
|
447 |
by (drule cardR_determ_aux1) auto |
|
448 |
||
449 |
lemma cardR_determ: "(A, m): cardR ==> (!!n. (A, n) : cardR ==> n = m)" |
|
450 |
apply (induct set: cardR) |
|
451 |
apply (safe elim!: cardR_emptyE cardR_insertE) |
|
452 |
apply (rename_tac B b m) |
|
453 |
apply (case_tac "a = b") |
|
454 |
apply (subgoal_tac "A = B") |
|
| 14208 | 455 |
prefer 2 apply (blast elim: equalityE, blast) |
| 12396 | 456 |
apply (subgoal_tac "EX C. A = insert b C & B = insert a C") |
457 |
prefer 2 |
|
458 |
apply (rule_tac x = "A Int B" in exI) |
|
459 |
apply (blast elim: equalityE) |
|
460 |
apply (frule_tac A = B in cardR_SucD) |
|
461 |
apply (blast intro!: cardR.intros dest!: cardR_determ_aux2) |
|
462 |
done |
|
463 |
||
464 |
lemma cardR_imp_finite: "(A, n) : cardR ==> finite A" |
|
465 |
by (induct set: cardR) simp_all |
|
466 |
||
467 |
lemma finite_imp_cardR: "finite A ==> EX n. (A, n) : cardR" |
|
468 |
by (induct set: Finites) (auto intro!: cardR.intros) |
|
469 |
||
470 |
lemma card_equality: "(A,n) : cardR ==> card A = n" |
|
471 |
by (unfold card_def) (blast intro: cardR_determ) |
|
472 |
||
473 |
lemma card_empty [simp]: "card {} = 0"
|
|
474 |
by (unfold card_def) (blast intro!: cardR.intros elim!: cardR_emptyE) |
|
475 |
||
476 |
lemma card_insert_disjoint [simp]: |
|
477 |
"finite A ==> x \<notin> A ==> card (insert x A) = Suc(card A)" |
|
478 |
proof - |
|
479 |
assume x: "x \<notin> A" |
|
480 |
hence aux: "!!n. ((insert x A, n) : cardR) = (EX m. (A, m) : cardR & n = Suc m)" |
|
481 |
apply (auto intro!: cardR.intros) |
|
482 |
apply (rule_tac A1 = A in finite_imp_cardR [THEN exE]) |
|
483 |
apply (force dest: cardR_imp_finite) |
|
484 |
apply (blast intro!: cardR.intros intro: cardR_determ) |
|
485 |
done |
|
486 |
assume "finite A" |
|
487 |
thus ?thesis |
|
488 |
apply (simp add: card_def aux) |
|
489 |
apply (rule the_equality) |
|
490 |
apply (auto intro: finite_imp_cardR |
|
491 |
cong: conj_cong simp: card_def [symmetric] card_equality) |
|
492 |
done |
|
493 |
qed |
|
494 |
||
495 |
lemma card_0_eq [simp]: "finite A ==> (card A = 0) = (A = {})"
|
|
496 |
apply auto |
|
| 14208 | 497 |
apply (drule_tac a = x in mk_disjoint_insert, clarify) |
498 |
apply (rotate_tac -1, auto) |
|
| 12396 | 499 |
done |
500 |
||
501 |
lemma card_insert_if: |
|
502 |
"finite A ==> card (insert x A) = (if x:A then card A else Suc(card(A)))" |
|
503 |
by (simp add: insert_absorb) |
|
504 |
||
505 |
lemma card_Suc_Diff1: "finite A ==> x: A ==> Suc (card (A - {x})) = card A"
|
|
| 14302 | 506 |
apply(rule_tac t = A in insert_Diff [THEN subst], assumption) |
507 |
apply(simp del:insert_Diff_single) |
|
508 |
done |
|
| 12396 | 509 |
|
510 |
lemma card_Diff_singleton: |
|
511 |
"finite A ==> x: A ==> card (A - {x}) = card A - 1"
|
|
512 |
by (simp add: card_Suc_Diff1 [symmetric]) |
|
513 |
||
514 |
lemma card_Diff_singleton_if: |
|
515 |
"finite A ==> card (A-{x}) = (if x : A then card A - 1 else card A)"
|
|
516 |
by (simp add: card_Diff_singleton) |
|
517 |
||
518 |
lemma card_insert: "finite A ==> card (insert x A) = Suc (card (A - {x}))"
|
|
519 |
by (simp add: card_insert_if card_Suc_Diff1) |
|
520 |
||
521 |
lemma card_insert_le: "finite A ==> card A <= card (insert x A)" |
|
522 |
by (simp add: card_insert_if) |
|
523 |
||
524 |
lemma card_seteq: "finite B ==> (!!A. A <= B ==> card B <= card A ==> A = B)" |
|
| 14208 | 525 |
apply (induct set: Finites, simp, clarify) |
| 12396 | 526 |
apply (subgoal_tac "finite A & A - {x} <= F")
|
| 14208 | 527 |
prefer 2 apply (blast intro: finite_subset, atomize) |
| 12396 | 528 |
apply (drule_tac x = "A - {x}" in spec)
|
529 |
apply (simp add: card_Diff_singleton_if split add: split_if_asm) |
|
| 14208 | 530 |
apply (case_tac "card A", auto) |
| 12396 | 531 |
done |
532 |
||
533 |
lemma psubset_card_mono: "finite B ==> A < B ==> card A < card B" |
|
534 |
apply (simp add: psubset_def linorder_not_le [symmetric]) |
|
535 |
apply (blast dest: card_seteq) |
|
536 |
done |
|
537 |
||
538 |
lemma card_mono: "finite B ==> A <= B ==> card A <= card B" |
|
| 14208 | 539 |
apply (case_tac "A = B", simp) |
| 12396 | 540 |
apply (simp add: linorder_not_less [symmetric]) |
541 |
apply (blast dest: card_seteq intro: order_less_imp_le) |
|
542 |
done |
|
543 |
||
544 |
lemma card_Un_Int: "finite A ==> finite B |
|
545 |
==> card A + card B = card (A Un B) + card (A Int B)" |
|
| 14208 | 546 |
apply (induct set: Finites, simp) |
| 12396 | 547 |
apply (simp add: insert_absorb Int_insert_left) |
548 |
done |
|
549 |
||
550 |
lemma card_Un_disjoint: "finite A ==> finite B |
|
551 |
==> A Int B = {} ==> card (A Un B) = card A + card B"
|
|
552 |
by (simp add: card_Un_Int) |
|
553 |
||
554 |
lemma card_Diff_subset: |
|
555 |
"finite A ==> B <= A ==> card A - card B = card (A - B)" |
|
556 |
apply (subgoal_tac "(A - B) Un B = A") |
|
557 |
prefer 2 apply blast |
|
| 14331 | 558 |
apply (rule nat_add_right_cancel [THEN iffD1]) |
| 12396 | 559 |
apply (rule card_Un_disjoint [THEN subst]) |
560 |
apply (erule_tac [4] ssubst) |
|
561 |
prefer 3 apply blast |
|
562 |
apply (simp_all add: add_commute not_less_iff_le |
|
563 |
add_diff_inverse card_mono finite_subset) |
|
564 |
done |
|
565 |
||
566 |
lemma card_Diff1_less: "finite A ==> x: A ==> card (A - {x}) < card A"
|
|
567 |
apply (rule Suc_less_SucD) |
|
568 |
apply (simp add: card_Suc_Diff1) |
|
569 |
done |
|
570 |
||
571 |
lemma card_Diff2_less: |
|
572 |
"finite A ==> x: A ==> y: A ==> card (A - {x} - {y}) < card A"
|
|
573 |
apply (case_tac "x = y") |
|
574 |
apply (simp add: card_Diff1_less) |
|
575 |
apply (rule less_trans) |
|
576 |
prefer 2 apply (auto intro!: card_Diff1_less) |
|
577 |
done |
|
578 |
||
579 |
lemma card_Diff1_le: "finite A ==> card (A - {x}) <= card A"
|
|
580 |
apply (case_tac "x : A") |
|
581 |
apply (simp_all add: card_Diff1_less less_imp_le) |
|
582 |
done |
|
583 |
||
584 |
lemma card_psubset: "finite B ==> A \<subseteq> B ==> card A < card B ==> A < B" |
|
| 14208 | 585 |
by (erule psubsetI, blast) |
| 12396 | 586 |
|
587 |
||
588 |
subsubsection {* Cardinality of image *}
|
|
589 |
||
590 |
lemma card_image_le: "finite A ==> card (f ` A) <= card A" |
|
| 14208 | 591 |
apply (induct set: Finites, simp) |
| 12396 | 592 |
apply (simp add: le_SucI finite_imageI card_insert_if) |
593 |
done |
|
594 |
||
595 |
lemma card_image: "finite A ==> inj_on f A ==> card (f ` A) = card A" |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
596 |
apply (induct set: Finites, simp_all, atomize, safe) |
| 14208 | 597 |
apply (unfold inj_on_def, blast) |
| 12396 | 598 |
apply (subst card_insert_disjoint) |
| 14208 | 599 |
apply (erule finite_imageI, blast, blast) |
| 12396 | 600 |
done |
601 |
||
602 |
lemma endo_inj_surj: "finite A ==> f ` A \<subseteq> A ==> inj_on f A ==> f ` A = A" |
|
603 |
by (simp add: card_seteq card_image) |
|
604 |
||
605 |
||
606 |
subsubsection {* Cardinality of the Powerset *}
|
|
607 |
||
608 |
lemma card_Pow: "finite A ==> card (Pow A) = Suc (Suc 0) ^ card A" (* FIXME numeral 2 (!?) *) |
|
609 |
apply (induct set: Finites) |
|
610 |
apply (simp_all add: Pow_insert) |
|
| 14208 | 611 |
apply (subst card_Un_disjoint, blast) |
612 |
apply (blast intro: finite_imageI, blast) |
|
| 12396 | 613 |
apply (subgoal_tac "inj_on (insert x) (Pow F)") |
614 |
apply (simp add: card_image Pow_insert) |
|
615 |
apply (unfold inj_on_def) |
|
616 |
apply (blast elim!: equalityE) |
|
617 |
done |
|
618 |
||
619 |
text {*
|
|
620 |
\medskip Relates to equivalence classes. Based on a theorem of |
|
621 |
F. Kammüller's. The @{prop "finite C"} premise is redundant.
|
|
622 |
*} |
|
623 |
||
624 |
lemma dvd_partition: |
|
625 |
"finite C ==> finite (Union C) ==> |
|
626 |
ALL c : C. k dvd card c ==> |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
627 |
(ALL c1: C. ALL c2: C. c1 \<noteq> c2 --> c1 Int c2 = {}) ==>
|
| 12396 | 628 |
k dvd card (Union C)" |
| 14208 | 629 |
apply (induct set: Finites, simp_all, clarify) |
| 12396 | 630 |
apply (subst card_Un_disjoint) |
631 |
apply (auto simp add: dvd_add disjoint_eq_subset_Compl) |
|
632 |
done |
|
633 |
||
634 |
||
635 |
subsection {* A fold functional for finite sets *}
|
|
636 |
||
637 |
text {*
|
|
638 |
For @{text n} non-negative we have @{text "fold f e {x1, ..., xn} =
|
|
639 |
f x1 (... (f xn e))"} where @{text f} is at least left-commutative.
|
|
640 |
*} |
|
641 |
||
642 |
consts |
|
643 |
foldSet :: "('b => 'a => 'a) => 'a => ('b set \<times> 'a) set"
|
|
644 |
||
645 |
inductive "foldSet f e" |
|
646 |
intros |
|
647 |
emptyI [intro]: "({}, e) : foldSet f e"
|
|
648 |
insertI [intro]: "x \<notin> A ==> (A, y) : foldSet f e ==> (insert x A, f x y) : foldSet f e" |
|
649 |
||
650 |
inductive_cases empty_foldSetE [elim!]: "({}, x) : foldSet f e"
|
|
651 |
||
652 |
constdefs |
|
653 |
fold :: "('b => 'a => 'a) => 'a => 'b set => 'a"
|
|
654 |
"fold f e A == THE x. (A, x) : foldSet f e" |
|
655 |
||
656 |
lemma Diff1_foldSet: "(A - {x}, y) : foldSet f e ==> x: A ==> (A, f x y) : foldSet f e"
|
|
| 14208 | 657 |
by (erule insert_Diff [THEN subst], rule foldSet.intros, auto) |
| 12396 | 658 |
|
659 |
lemma foldSet_imp_finite [simp]: "(A, x) : foldSet f e ==> finite A" |
|
660 |
by (induct set: foldSet) auto |
|
661 |
||
662 |
lemma finite_imp_foldSet: "finite A ==> EX x. (A, x) : foldSet f e" |
|
663 |
by (induct set: Finites) auto |
|
664 |
||
665 |
||
666 |
subsubsection {* Left-commutative operations *}
|
|
667 |
||
668 |
locale LC = |
|
669 |
fixes f :: "'b => 'a => 'a" (infixl "\<cdot>" 70) |
|
670 |
assumes left_commute: "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
|
671 |
||
672 |
lemma (in LC) foldSet_determ_aux: |
|
673 |
"ALL A x. card A < n --> (A, x) : foldSet f e --> |
|
674 |
(ALL y. (A, y) : foldSet f e --> y = x)" |
|
675 |
apply (induct n) |
|
676 |
apply (auto simp add: less_Suc_eq) |
|
| 14208 | 677 |
apply (erule foldSet.cases, blast) |
678 |
apply (erule foldSet.cases, blast, clarify) |
|
| 12396 | 679 |
txt {* force simplification of @{text "card A < card (insert ...)"}. *}
|
680 |
apply (erule rev_mp) |
|
681 |
apply (simp add: less_Suc_eq_le) |
|
682 |
apply (rule impI) |
|
683 |
apply (rename_tac Aa xa ya Ab xb yb, case_tac "xa = xb") |
|
684 |
apply (subgoal_tac "Aa = Ab") |
|
| 14208 | 685 |
prefer 2 apply (blast elim!: equalityE, blast) |
| 12396 | 686 |
txt {* case @{prop "xa \<notin> xb"}. *}
|
687 |
apply (subgoal_tac "Aa - {xb} = Ab - {xa} & xb : Aa & xa : Ab")
|
|
| 14208 | 688 |
prefer 2 apply (blast elim!: equalityE, clarify) |
| 12396 | 689 |
apply (subgoal_tac "Aa = insert xb Ab - {xa}")
|
690 |
prefer 2 apply blast |
|
691 |
apply (subgoal_tac "card Aa <= card Ab") |
|
692 |
prefer 2 |
|
693 |
apply (rule Suc_le_mono [THEN subst]) |
|
694 |
apply (simp add: card_Suc_Diff1) |
|
695 |
apply (rule_tac A1 = "Aa - {xb}" and f1 = f in finite_imp_foldSet [THEN exE])
|
|
696 |
apply (blast intro: foldSet_imp_finite finite_Diff) |
|
697 |
apply (frule (1) Diff1_foldSet) |
|
698 |
apply (subgoal_tac "ya = f xb x") |
|
699 |
prefer 2 apply (blast del: equalityCE) |
|
700 |
apply (subgoal_tac "(Ab - {xa}, x) : foldSet f e")
|
|
701 |
prefer 2 apply simp |
|
702 |
apply (subgoal_tac "yb = f xa x") |
|
703 |
prefer 2 apply (blast del: equalityCE dest: Diff1_foldSet) |
|
704 |
apply (simp (no_asm_simp) add: left_commute) |
|
705 |
done |
|
706 |
||
707 |
lemma (in LC) foldSet_determ: "(A, x) : foldSet f e ==> (A, y) : foldSet f e ==> y = x" |
|
708 |
by (blast intro: foldSet_determ_aux [rule_format]) |
|
709 |
||
710 |
lemma (in LC) fold_equality: "(A, y) : foldSet f e ==> fold f e A = y" |
|
711 |
by (unfold fold_def) (blast intro: foldSet_determ) |
|
712 |
||
713 |
lemma fold_empty [simp]: "fold f e {} = e"
|
|
714 |
by (unfold fold_def) blast |
|
715 |
||
716 |
lemma (in LC) fold_insert_aux: "x \<notin> A ==> |
|
717 |
((insert x A, v) : foldSet f e) = |
|
718 |
(EX y. (A, y) : foldSet f e & v = f x y)" |
|
719 |
apply auto |
|
720 |
apply (rule_tac A1 = A and f1 = f in finite_imp_foldSet [THEN exE]) |
|
721 |
apply (fastsimp dest: foldSet_imp_finite) |
|
722 |
apply (blast intro: foldSet_determ) |
|
723 |
done |
|
724 |
||
725 |
lemma (in LC) fold_insert: |
|
726 |
"finite A ==> x \<notin> A ==> fold f e (insert x A) = f x (fold f e A)" |
|
727 |
apply (unfold fold_def) |
|
728 |
apply (simp add: fold_insert_aux) |
|
729 |
apply (rule the_equality) |
|
730 |
apply (auto intro: finite_imp_foldSet |
|
731 |
cong add: conj_cong simp add: fold_def [symmetric] fold_equality) |
|
732 |
done |
|
733 |
||
734 |
lemma (in LC) fold_commute: "finite A ==> (!!e. f x (fold f e A) = fold f (f x e) A)" |
|
| 14208 | 735 |
apply (induct set: Finites, simp) |
| 12396 | 736 |
apply (simp add: left_commute fold_insert) |
737 |
done |
|
738 |
||
739 |
lemma (in LC) fold_nest_Un_Int: |
|
740 |
"finite A ==> finite B |
|
741 |
==> fold f (fold f e B) A = fold f (fold f e (A Int B)) (A Un B)" |
|
| 14208 | 742 |
apply (induct set: Finites, simp) |
| 12396 | 743 |
apply (simp add: fold_insert fold_commute Int_insert_left insert_absorb) |
744 |
done |
|
745 |
||
746 |
lemma (in LC) fold_nest_Un_disjoint: |
|
747 |
"finite A ==> finite B ==> A Int B = {}
|
|
748 |
==> fold f e (A Un B) = fold f (fold f e B) A" |
|
749 |
by (simp add: fold_nest_Un_Int) |
|
750 |
||
751 |
declare foldSet_imp_finite [simp del] |
|
752 |
empty_foldSetE [rule del] foldSet.intros [rule del] |
|
753 |
-- {* Delete rules to do with @{text foldSet} relation. *}
|
|
754 |
||
755 |
||
756 |
||
757 |
subsubsection {* Commutative monoids *}
|
|
758 |
||
759 |
text {*
|
|
760 |
We enter a more restrictive context, with @{text "f :: 'a => 'a => 'a"}
|
|
761 |
instead of @{text "'b => 'a => 'a"}.
|
|
762 |
*} |
|
763 |
||
764 |
locale ACe = |
|
765 |
fixes f :: "'a => 'a => 'a" (infixl "\<cdot>" 70) |
|
766 |
and e :: 'a |
|
767 |
assumes ident [simp]: "x \<cdot> e = x" |
|
768 |
and commute: "x \<cdot> y = y \<cdot> x" |
|
769 |
and assoc: "(x \<cdot> y) \<cdot> z = x \<cdot> (y \<cdot> z)" |
|
770 |
||
771 |
lemma (in ACe) left_commute: "x \<cdot> (y \<cdot> z) = y \<cdot> (x \<cdot> z)" |
|
772 |
proof - |
|
773 |
have "x \<cdot> (y \<cdot> z) = (y \<cdot> z) \<cdot> x" by (simp only: commute) |
|
774 |
also have "... = y \<cdot> (z \<cdot> x)" by (simp only: assoc) |
|
775 |
also have "z \<cdot> x = x \<cdot> z" by (simp only: commute) |
|
776 |
finally show ?thesis . |
|
777 |
qed |
|
778 |
||
| 12718 | 779 |
lemmas (in ACe) AC = assoc commute left_commute |
| 12396 | 780 |
|
| 12693 | 781 |
lemma (in ACe) left_ident [simp]: "e \<cdot> x = x" |
| 12396 | 782 |
proof - |
783 |
have "x \<cdot> e = x" by (rule ident) |
|
784 |
thus ?thesis by (subst commute) |
|
785 |
qed |
|
786 |
||
787 |
lemma (in ACe) fold_Un_Int: |
|
788 |
"finite A ==> finite B ==> |
|
789 |
fold f e A \<cdot> fold f e B = fold f e (A Un B) \<cdot> fold f e (A Int B)" |
|
| 14208 | 790 |
apply (induct set: Finites, simp) |
| 13400 | 791 |
apply (simp add: AC insert_absorb Int_insert_left |
| 13421 | 792 |
LC.fold_insert [OF LC.intro]) |
| 12396 | 793 |
done |
794 |
||
795 |
lemma (in ACe) fold_Un_disjoint: |
|
796 |
"finite A ==> finite B ==> A Int B = {} ==>
|
|
797 |
fold f e (A Un B) = fold f e A \<cdot> fold f e B" |
|
798 |
by (simp add: fold_Un_Int) |
|
799 |
||
800 |
lemma (in ACe) fold_Un_disjoint2: |
|
801 |
"finite A ==> finite B ==> A Int B = {} ==>
|
|
802 |
fold (f o g) e (A Un B) = fold (f o g) e A \<cdot> fold (f o g) e B" |
|
803 |
proof - |
|
804 |
assume b: "finite B" |
|
805 |
assume "finite A" |
|
806 |
thus "A Int B = {} ==>
|
|
807 |
fold (f o g) e (A Un B) = fold (f o g) e A \<cdot> fold (f o g) e B" |
|
808 |
proof induct |
|
809 |
case empty |
|
810 |
thus ?case by simp |
|
811 |
next |
|
812 |
case (insert F x) |
|
| 13571 | 813 |
have "fold (f o g) e (insert x F \<union> B) = fold (f o g) e (insert x (F \<union> B))" |
| 12396 | 814 |
by simp |
| 13571 | 815 |
also have "... = (f o g) x (fold (f o g) e (F \<union> B))" |
| 13400 | 816 |
by (rule LC.fold_insert [OF LC.intro]) |
| 13421 | 817 |
(insert b insert, auto simp add: left_commute) |
| 13571 | 818 |
also from insert have "fold (f o g) e (F \<union> B) = |
819 |
fold (f o g) e F \<cdot> fold (f o g) e B" by blast |
|
820 |
also have "(f o g) x ... = (f o g) x (fold (f o g) e F) \<cdot> fold (f o g) e B" |
|
| 12396 | 821 |
by (simp add: AC) |
| 13571 | 822 |
also have "(f o g) x (fold (f o g) e F) = fold (f o g) e (insert x F)" |
| 13400 | 823 |
by (rule LC.fold_insert [OF LC.intro, symmetric]) (insert b insert, |
| 13421 | 824 |
auto simp add: left_commute) |
| 12396 | 825 |
finally show ?case . |
826 |
qed |
|
827 |
qed |
|
828 |
||
829 |
||
830 |
subsection {* Generalized summation over a set *}
|
|
831 |
||
832 |
constdefs |
|
833 |
setsum :: "('a => 'b) => 'a set => 'b::plus_ac0"
|
|
834 |
"setsum f A == if finite A then fold (op + o f) 0 A else 0" |
|
835 |
||
836 |
syntax |
|
837 |
"_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0" ("\<Sum>_:_. _" [0, 51, 10] 10)
|
|
838 |
syntax (xsymbols) |
|
839 |
"_setsum" :: "idt => 'a set => 'b => 'b::plus_ac0" ("\<Sum>_\<in>_. _" [0, 51, 10] 10)
|
|
840 |
translations |
|
841 |
"\<Sum>i:A. b" == "setsum (%i. b) A" -- {* Beware of argument permutation! *}
|
|
842 |
||
843 |
||
844 |
lemma setsum_empty [simp]: "setsum f {} = 0"
|
|
845 |
by (simp add: setsum_def) |
|
846 |
||
847 |
lemma setsum_insert [simp]: |
|
848 |
"finite F ==> a \<notin> F ==> setsum f (insert a F) = f a + setsum f F" |
|
| 13365 | 849 |
by (simp add: setsum_def |
| 13421 | 850 |
LC.fold_insert [OF LC.intro] plus_ac0_left_commute) |
| 12396 | 851 |
|
852 |
lemma setsum_0: "setsum (\<lambda>i. 0) A = 0" |
|
853 |
apply (case_tac "finite A") |
|
854 |
prefer 2 apply (simp add: setsum_def) |
|
| 14208 | 855 |
apply (erule finite_induct, auto) |
| 12396 | 856 |
done |
857 |
||
858 |
lemma setsum_SucD: "setsum f A = Suc n ==> EX a:A. 0 < f a" |
|
859 |
apply (case_tac "finite A") |
|
860 |
prefer 2 apply (simp add: setsum_def) |
|
861 |
apply (erule rev_mp) |
|
| 14208 | 862 |
apply (erule finite_induct, auto) |
| 12396 | 863 |
done |
864 |
||
865 |
lemma card_eq_setsum: "finite A ==> card A = setsum (\<lambda>x. 1) A" |
|
866 |
-- {* Could allow many @{text "card"} proofs to be simplified. *}
|
|
867 |
by (induct set: Finites) auto |
|
868 |
||
869 |
lemma setsum_Un_Int: "finite A ==> finite B |
|
870 |
==> setsum g (A Un B) + setsum g (A Int B) = setsum g A + setsum g B" |
|
871 |
-- {* The reversed orientation looks more natural, but LOOPS as a simprule! *}
|
|
| 14208 | 872 |
apply (induct set: Finites, simp) |
| 12396 | 873 |
apply (simp add: plus_ac0 Int_insert_left insert_absorb) |
874 |
done |
|
875 |
||
876 |
lemma setsum_Un_disjoint: "finite A ==> finite B |
|
877 |
==> A Int B = {} ==> setsum g (A Un B) = setsum g A + setsum g B"
|
|
| 14208 | 878 |
apply (subst setsum_Un_Int [symmetric], auto) |
| 12396 | 879 |
done |
880 |
||
|
12937
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
881 |
lemma setsum_UN_disjoint: |
|
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
882 |
fixes f :: "'a => 'b::plus_ac0" |
|
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
883 |
shows |
|
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
884 |
"finite I ==> (ALL i:I. finite (A i)) ==> |
|
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
885 |
(ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
|
|
0c4fd7529467
clarified syntax of ``long'' statements: fixes/assumes/shows;
wenzelm
parents:
12718
diff
changeset
|
886 |
setsum f (UNION I A) = setsum (\<lambda>i. setsum f (A i)) I" |
| 14208 | 887 |
apply (induct set: Finites, simp, atomize) |
| 12396 | 888 |
apply (subgoal_tac "ALL i:F. x \<noteq> i") |
889 |
prefer 2 apply blast |
|
890 |
apply (subgoal_tac "A x Int UNION F A = {}")
|
|
891 |
prefer 2 apply blast |
|
892 |
apply (simp add: setsum_Un_disjoint) |
|
893 |
done |
|
894 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
895 |
lemma setsum_Union_disjoint: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
896 |
"finite C ==> (ALL A:C. finite A) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
897 |
(ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
898 |
setsum f (Union C) = setsum (setsum f) C" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
899 |
apply (frule setsum_UN_disjoint [of C id f]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
900 |
apply (unfold Union_def id_def, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
901 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
902 |
|
| 12396 | 903 |
lemma setsum_addf: "setsum (\<lambda>x. f x + g x) A = (setsum f A + setsum g A)" |
904 |
apply (case_tac "finite A") |
|
905 |
prefer 2 apply (simp add: setsum_def) |
|
| 14208 | 906 |
apply (erule finite_induct, auto) |
| 12396 | 907 |
apply (simp add: plus_ac0) |
908 |
done |
|
909 |
||
910 |
lemma setsum_cong: |
|
911 |
"A = B ==> (!!x. x:B ==> f x = g x) ==> setsum f A = setsum g B" |
|
912 |
apply (case_tac "finite B") |
|
| 14208 | 913 |
prefer 2 apply (simp add: setsum_def, simp) |
| 12396 | 914 |
apply (subgoal_tac "ALL C. C <= B --> (ALL x:C. f x = g x) --> setsum f C = setsum g C") |
915 |
apply simp |
|
| 14208 | 916 |
apply (erule finite_induct, simp) |
917 |
apply (simp add: subset_insert_iff, clarify) |
|
| 12396 | 918 |
apply (subgoal_tac "finite C") |
919 |
prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl]) |
|
920 |
apply (subgoal_tac "C = insert x (C - {x})")
|
|
921 |
prefer 2 apply blast |
|
922 |
apply (erule ssubst) |
|
923 |
apply (drule spec) |
|
924 |
apply (erule (1) notE impE) |
|
| 14302 | 925 |
apply (simp add: Ball_def del:insert_Diff_single) |
| 12396 | 926 |
done |
927 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
928 |
lemma card_UN_disjoint: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
929 |
fixes f :: "'a => 'b::plus_ac0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
930 |
shows |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
931 |
"finite I ==> (ALL i:I. finite (A i)) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
932 |
(ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
933 |
card (UNION I A) = setsum (\<lambda>i. card (A i)) I" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
934 |
apply (subst card_eq_setsum) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
935 |
apply (subst finite_UN, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
936 |
apply (subgoal_tac "setsum (\<lambda>i. card (A i)) I = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
937 |
setsum (%i. (setsum (%x. 1) (A i))) I") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
938 |
apply (erule ssubst) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
939 |
apply (erule setsum_UN_disjoint, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
940 |
apply (rule setsum_cong) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
941 |
apply (simp, simp add: card_eq_setsum) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
942 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
943 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
944 |
lemma card_Union_disjoint: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
945 |
"finite C ==> (ALL A:C. finite A) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
946 |
(ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
947 |
card (Union C) = setsum card C" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
948 |
apply (frule card_UN_disjoint [of C id]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
949 |
apply (unfold Union_def id_def, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
950 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
951 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
952 |
lemma setsum_0': "ALL a:F. f a = 0 ==> setsum f F = 0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
953 |
apply (subgoal_tac "setsum f F = setsum (%x. 0) F") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
954 |
apply (erule ssubst, rule setsum_0) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
955 |
apply (rule setsum_cong, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
956 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
957 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
958 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
959 |
subsubsection {* Reindexing sums *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
960 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
961 |
lemma setsum_reindex [rule_format]: "finite B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
962 |
inj_on f B --> setsum h (f ` B) = setsum (h \<circ> f) B" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
963 |
apply (rule finite_induct, assumption, force) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
964 |
apply (rule impI, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
965 |
apply (simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
966 |
apply (subgoal_tac "f x \<notin> f ` F") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
967 |
apply (subgoal_tac "finite (f ` F)") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
968 |
apply (auto simp add: setsum_insert) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
969 |
apply (simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
970 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
971 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
972 |
lemma setsum_reindex_id: "finite B ==> inj_on f B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
973 |
setsum f B = setsum id (f ` B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
974 |
by (auto simp add: setsum_reindex id_o) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
975 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
976 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
977 |
subsubsection {* Properties in more restricted classes of structures *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
978 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
979 |
lemma setsum_eq_0_iff [simp]: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
980 |
"finite F ==> (setsum f F = 0) = (ALL a:F. f a = (0::nat))" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
981 |
by (induct set: Finites) auto |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
982 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
983 |
lemma setsum_constant_nat: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
984 |
"finite A ==> (\<Sum>x: A. y) = (card A) * y" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
985 |
-- {* Later generalized to any semiring. *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
986 |
by (erule finite_induct, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
987 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
988 |
lemma setsum_Un: "finite A ==> finite B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
989 |
(setsum f (A Un B) :: nat) = setsum f A + setsum f B - setsum f (A Int B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
990 |
-- {* For the natural numbers, we have subtraction. *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
991 |
by (subst setsum_Un_Int [symmetric], auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
992 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
993 |
lemma setsum_Un_ring: "finite A ==> finite B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
994 |
(setsum f (A Un B) :: 'a :: ring) = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
995 |
setsum f A + setsum f B - setsum f (A Int B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
996 |
apply (subst setsum_Un_Int [symmetric], auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
997 |
apply (simp add: compare_rls) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
998 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
999 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1000 |
lemma setsum_diff1: "(setsum f (A - {a}) :: nat) =
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1001 |
(if a:A then setsum f A - f a else setsum f A)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1002 |
apply (case_tac "finite A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1003 |
prefer 2 apply (simp add: setsum_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1004 |
apply (erule finite_induct) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1005 |
apply (auto simp add: insert_Diff_if) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1006 |
apply (drule_tac a = a in mk_disjoint_insert, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1007 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1008 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1009 |
lemma setsum_negf: "finite A ==> setsum (%x. - (f x)::'a::ring) A = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1010 |
- setsum f A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1011 |
by (induct set: Finites, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1012 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1013 |
lemma setsum_subtractf: "finite A ==> setsum (%x. ((f x)::'a::ring) - g x) A = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1014 |
setsum f A - setsum g A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1015 |
by (simp add: diff_minus setsum_addf setsum_negf) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1016 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1017 |
lemma setsum_nonneg: "[| finite A; |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1018 |
\<forall>x \<in> A. (0::'a::ordered_semiring) \<le> f x |] ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1019 |
0 \<le> setsum f A"; |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1020 |
apply (induct set: Finites, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1021 |
apply (subgoal_tac "0 + 0 \<le> f x + setsum f F", simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1022 |
apply (blast intro: add_mono) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1023 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1024 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1025 |
subsubsection {* Cardinality of Sigma and Cartesian product *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1026 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1027 |
lemma SigmaI_insert: "y \<notin> A ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1028 |
(SIGMA x:(insert y A). B x) = (({y} <*> (B y)) \<union> (SIGMA x: A. B x))"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1029 |
by auto |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1030 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1031 |
lemma card_cartesian_product_singleton [simp]: "finite A ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1032 |
card({x} <*> A) = card(A)"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1033 |
apply (subgoal_tac "inj_on (%y .(x,y)) A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1034 |
apply (frule card_image, assumption) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1035 |
apply (subgoal_tac "(Pair x ` A) = {x} <*> A")
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1036 |
apply (auto simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1037 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1038 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1039 |
lemma card_SigmaI [rule_format,simp]: "finite A ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1040 |
(ALL a:A. finite (B a)) --> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1041 |
card (SIGMA x: A. B x) = (\<Sum>a: A. card (B a))" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1042 |
apply (erule finite_induct, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1043 |
apply (subst SigmaI_insert, assumption) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1044 |
apply (subst card_Un_disjoint) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1045 |
apply (auto intro: finite_SigmaI) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1046 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1047 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1048 |
lemma card_cartesian_product [simp]: "[| finite A; finite B |] ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1049 |
card (A <*> B) = card(A) * card(B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1050 |
apply (subst card_SigmaI, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1051 |
apply (simp add: setsum_constant_nat) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1052 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1053 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1054 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1055 |
subsection {* Generalized product over a set *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1056 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1057 |
constdefs |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1058 |
setprod :: "('a => 'b) => 'a set => 'b::times_ac1"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1059 |
"setprod f A == if finite A then fold (op * o f) 1 A else 1" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1060 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1061 |
syntax |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1062 |
"_setprod" :: "idt => 'a set => 'b => 'b::plus_ac0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1063 |
("\<Prod>_:_. _" [0, 51, 10] 10)
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1064 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1065 |
syntax (xsymbols) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1066 |
"_setprod" :: "idt => 'a set => 'b => 'b::plus_ac0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1067 |
("\<Prod>_\<in>_. _" [0, 51, 10] 10)
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1068 |
translations |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1069 |
"\<Prod>i:A. b" == "setprod (%i. b) A" -- {* Beware of argument permutation! *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1070 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1071 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1072 |
lemma setprod_empty [simp]: "setprod f {} = 1"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1073 |
by (auto simp add: setprod_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1074 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1075 |
lemma setprod_insert [simp]: "[| finite A; a \<notin> A |] ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1076 |
setprod f (insert a A) = f a * setprod f A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1077 |
by (auto simp add: setprod_def LC_def LC.fold_insert |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1078 |
times_ac1_left_commute) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1079 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1080 |
lemma setprod_1: "setprod (\<lambda>i. 1) A = 1" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1081 |
apply (case_tac "finite A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1082 |
apply (erule finite_induct, auto simp add: times_ac1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1083 |
apply (simp add: setprod_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1084 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1085 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1086 |
lemma setprod_Un_Int: "finite A ==> finite B |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1087 |
==> setprod g (A Un B) * setprod g (A Int B) = setprod g A * setprod g B" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1088 |
apply (induct set: Finites, simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1089 |
apply (simp add: times_ac1 insert_absorb) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1090 |
apply (simp add: times_ac1 Int_insert_left insert_absorb) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1091 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1092 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1093 |
lemma setprod_Un_disjoint: "finite A ==> finite B |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1094 |
==> A Int B = {} ==> setprod g (A Un B) = setprod g A * setprod g B"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1095 |
apply (subst setprod_Un_Int [symmetric], auto simp add: times_ac1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1096 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1097 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1098 |
lemma setprod_UN_disjoint: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1099 |
fixes f :: "'a => 'b::times_ac1" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1100 |
shows |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1101 |
"finite I ==> (ALL i:I. finite (A i)) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1102 |
(ALL i:I. ALL j:I. i \<noteq> j --> A i Int A j = {}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1103 |
setprod f (UNION I A) = setprod (\<lambda>i. setprod f (A i)) I" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1104 |
apply (induct set: Finites, simp, atomize) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1105 |
apply (subgoal_tac "ALL i:F. x \<noteq> i") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1106 |
prefer 2 apply blast |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1107 |
apply (subgoal_tac "A x Int UNION F A = {}")
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1108 |
prefer 2 apply blast |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1109 |
apply (simp add: setprod_Un_disjoint) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1110 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1111 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1112 |
lemma setprod_Union_disjoint: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1113 |
"finite C ==> (ALL A:C. finite A) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1114 |
(ALL A:C. ALL B:C. A \<noteq> B --> A Int B = {}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1115 |
setprod f (Union C) = setprod (setprod f) C" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1116 |
apply (frule setprod_UN_disjoint [of C id f]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1117 |
apply (unfold Union_def id_def, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1118 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1119 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1120 |
lemma setprod_timesf: "setprod (\<lambda>x. f x * g x) A = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1121 |
(setprod f A * setprod g A)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1122 |
apply (case_tac "finite A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1123 |
prefer 2 apply (simp add: setprod_def times_ac1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1124 |
apply (erule finite_induct, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1125 |
apply (simp add: times_ac1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1126 |
apply (simp add: times_ac1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1127 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1128 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1129 |
lemma setprod_cong: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1130 |
"A = B ==> (!!x. x:B ==> f x = g x) ==> setprod f A = setprod g B" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1131 |
apply (case_tac "finite B") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1132 |
prefer 2 apply (simp add: setprod_def, simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1133 |
apply (subgoal_tac "ALL C. C <= B --> (ALL x:C. f x = g x) --> setprod f C = setprod g C") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1134 |
apply simp |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1135 |
apply (erule finite_induct, simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1136 |
apply (simp add: subset_insert_iff, clarify) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1137 |
apply (subgoal_tac "finite C") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1138 |
prefer 2 apply (blast dest: finite_subset [COMP swap_prems_rl]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1139 |
apply (subgoal_tac "C = insert x (C - {x})")
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1140 |
prefer 2 apply blast |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1141 |
apply (erule ssubst) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1142 |
apply (drule spec) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1143 |
apply (erule (1) notE impE) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1144 |
apply (simp add: Ball_def del:insert_Diff_single) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1145 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1146 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1147 |
lemma setprod_1': "ALL a:F. f a = 1 ==> setprod f F = 1" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1148 |
apply (subgoal_tac "setprod f F = setprod (%x. 1) F") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1149 |
apply (erule ssubst, rule setprod_1) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1150 |
apply (rule setprod_cong, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1151 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1152 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1153 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1154 |
subsubsection {* Reindexing products *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1155 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1156 |
lemma setprod_reindex [rule_format]: "finite B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1157 |
inj_on f B --> setprod h (f ` B) = setprod (h \<circ> f) B" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1158 |
apply (rule finite_induct, assumption, force) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1159 |
apply (rule impI, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1160 |
apply (simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1161 |
apply (subgoal_tac "f x \<notin> f ` F") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1162 |
apply (subgoal_tac "finite (f ` F)") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1163 |
apply (auto simp add: setprod_insert) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1164 |
apply (simp add: inj_on_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1165 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1166 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1167 |
lemma setprod_reindex_id: "finite B ==> inj_on f B ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1168 |
setprod f B = setprod id (f ` B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1169 |
by (auto simp add: setprod_reindex id_o) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1170 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1171 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1172 |
subsubsection {* Properties in more restricted classes of structures *}
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1173 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1174 |
lemma setprod_eq_1_iff [simp]: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1175 |
"finite F ==> (setprod f F = 1) = (ALL a:F. f a = (1::nat))" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1176 |
by (induct set: Finites) auto |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1177 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1178 |
lemma setprod_constant: "finite A ==> (\<Prod>x: A. (y::'a::ringpower)) = |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1179 |
y^(card A)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1180 |
apply (erule finite_induct) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1181 |
apply (auto simp add: power_Suc) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1182 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1183 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1184 |
lemma setprod_zero: "finite A ==> EX x: A. f x = (0::'a::semiring) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1185 |
setprod f A = 0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1186 |
apply (induct set: Finites, force, clarsimp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1187 |
apply (erule disjE, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1188 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1189 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1190 |
lemma setprod_nonneg [rule_format]: "(ALL x: A. (0::'a::ordered_ring) \<le> f x) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1191 |
--> 0 \<le> setprod f A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1192 |
apply (case_tac "finite A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1193 |
apply (induct set: Finites, force, clarsimp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1194 |
apply (subgoal_tac "0 * 0 \<le> f x * setprod f F", force) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1195 |
apply (rule mult_mono, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1196 |
apply (auto simp add: setprod_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1197 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1198 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1199 |
lemma setprod_pos [rule_format]: "(ALL x: A. (0::'a::ordered_ring) < f x) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1200 |
--> 0 < setprod f A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1201 |
apply (case_tac "finite A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1202 |
apply (induct set: Finites, force, clarsimp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1203 |
apply (subgoal_tac "0 * 0 < f x * setprod f F", force) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1204 |
apply (rule mult_strict_mono, assumption+) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1205 |
apply (auto simp add: setprod_def) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1206 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1207 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1208 |
lemma setprod_nonzero [rule_format]: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1209 |
"(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1210 |
finite A ==> (ALL x: A. f x \<noteq> (0::'a)) --> setprod f A \<noteq> 0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1211 |
apply (erule finite_induct, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1212 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1213 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1214 |
lemma setprod_zero_eq: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1215 |
"(ALL x y. (x::'a::semiring) * y = 0 --> x = 0 | y = 0) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1216 |
finite A ==> (setprod f A = (0::'a)) = (EX x: A. f x = 0)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1217 |
apply (insert setprod_zero [of A f] setprod_nonzero [of A f], blast) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1218 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1219 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1220 |
lemma setprod_nonzero_field: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1221 |
"finite A ==> (ALL x: A. f x \<noteq> (0::'a::field)) ==> setprod f A \<noteq> 0" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1222 |
apply (rule setprod_nonzero, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1223 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1224 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1225 |
lemma setprod_zero_eq_field: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1226 |
"finite A ==> (setprod f A = (0::'a::field)) = (EX x: A. f x = 0)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1227 |
apply (rule setprod_zero_eq, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1228 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1229 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1230 |
lemma setprod_Un: "finite A ==> finite B ==> (ALL x: A Int B. f x \<noteq> 0) ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1231 |
(setprod f (A Un B) :: 'a ::{field})
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1232 |
= setprod f A * setprod f B / setprod f (A Int B)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1233 |
apply (subst setprod_Un_Int [symmetric], auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1234 |
apply (subgoal_tac "finite (A Int B)") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1235 |
apply (frule setprod_nonzero_field [of "A Int B" f], assumption) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1236 |
apply (subst times_divide_eq_right [THEN sym], auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1237 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1238 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1239 |
lemma setprod_diff1: "finite A ==> f a \<noteq> 0 ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1240 |
(setprod f (A - {a}) :: 'a :: {field}) =
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1241 |
(if a:A then setprod f A / f a else setprod f A)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1242 |
apply (erule finite_induct) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1243 |
apply (auto simp add: insert_Diff_if) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1244 |
apply (subgoal_tac "f a * setprod f F / f a = setprod f F * f a / f a") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1245 |
apply (erule ssubst) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1246 |
apply (subst times_divide_eq_right [THEN sym]) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1247 |
apply (auto simp add: mult_ac) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1248 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1249 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1250 |
lemma setprod_inversef: "finite A ==> |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1251 |
ALL x: A. f x \<noteq> (0::'a::{field,division_by_zero}) ==>
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1252 |
setprod (inverse \<circ> f) A = inverse (setprod f A)" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1253 |
apply (erule finite_induct) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1254 |
apply (simp, simp) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1255 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1256 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1257 |
lemma setprod_dividef: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1258 |
"[|finite A; |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1259 |
\<forall>x \<in> A. g x \<noteq> (0::'a::{field,division_by_zero})|]
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1260 |
==> setprod (%x. f x / g x) A = setprod f A / setprod g A" |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1261 |
apply (subgoal_tac |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1262 |
"setprod (%x. f x / g x) A = setprod (%x. f x * (inverse \<circ> g) x) A") |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1263 |
apply (erule ssubst) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1264 |
apply (subst divide_inverse) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1265 |
apply (subst setprod_timesf) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1266 |
apply (subst setprod_inversef, assumption+, rule refl) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1267 |
apply (rule setprod_cong, rule refl) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1268 |
apply (subst divide_inverse, auto) |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1269 |
done |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1270 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1271 |
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1272 |
subsection{* Min and Max of finite linearly ordered sets *}
|
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1273 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1274 |
text{* Seemed easier to define directly than via fold. *}
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1275 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1276 |
lemma ex_Max: fixes S :: "('a::linorder)set"
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1277 |
assumes fin: "finite S" shows "S \<noteq> {} ==> \<exists>m\<in>S. \<forall>s \<in> S. s \<le> m"
|
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1278 |
using fin |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1279 |
proof (induct) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1280 |
case empty thus ?case by simp |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1281 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1282 |
case (insert S x) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1283 |
show ?case |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1284 |
proof (cases) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1285 |
assume "S = {}" thus ?thesis by simp
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1286 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1287 |
assume nonempty: "S \<noteq> {}"
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1288 |
then obtain m where m: "m\<in>S" "\<forall>s\<in>S. s \<le> m" using insert by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1289 |
show ?thesis |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1290 |
proof (cases) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1291 |
assume "x \<le> m" thus ?thesis using m by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1292 |
next |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1293 |
assume "~ x \<le> m" thus ?thesis using m |
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1294 |
by(simp add:linorder_not_le order_less_le)(blast intro: order_trans) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1295 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1296 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1297 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1298 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1299 |
lemma ex_Min: fixes S :: "('a::linorder)set"
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1300 |
assumes fin: "finite S" shows "S \<noteq> {} ==> \<exists>m\<in>S. \<forall>s \<in> S. m \<le> s"
|
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1301 |
using fin |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1302 |
proof (induct) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1303 |
case empty thus ?case by simp |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1304 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1305 |
case (insert S x) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1306 |
show ?case |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1307 |
proof (cases) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1308 |
assume "S = {}" thus ?thesis by simp
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1309 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1310 |
assume nonempty: "S \<noteq> {}"
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1311 |
then obtain m where m: "m\<in>S" "\<forall>s\<in>S. m \<le> s" using insert by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1312 |
show ?thesis |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1313 |
proof (cases) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1314 |
assume "m \<le> x" thus ?thesis using m by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1315 |
next |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1316 |
assume "~ m \<le> x" thus ?thesis using m |
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1317 |
by(simp add:linorder_not_le order_less_le)(blast intro: order_trans) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1318 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1319 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1320 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1321 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1322 |
constdefs |
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1323 |
Min :: "('a::linorder)set => 'a"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1324 |
"Min S == THE m. m \<in> S \<and> (\<forall>s \<in> S. m \<le> s)" |
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1325 |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1326 |
Max :: "('a::linorder)set => 'a"
|
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1327 |
"Max S == THE m. m \<in> S \<and> (\<forall>s \<in> S. s \<le> m)" |
|
13490
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1328 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1329 |
lemma Min[simp]: assumes a: "finite S" "S \<noteq> {}"
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1330 |
shows "Min S \<in> S \<and> (\<forall>s \<in> S. Min S \<le> s)" (is "?P(Min S)") |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1331 |
proof (unfold Min_def, rule theI') |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1332 |
show "\<exists>!m. ?P m" |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1333 |
proof (rule ex_ex1I) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1334 |
show "\<exists>m. ?P m" using ex_Min[OF a] by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1335 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1336 |
fix m1 m2 assume "?P m1" "?P m2" |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1337 |
thus "m1 = m2" by (blast dest:order_antisym) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1338 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1339 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1340 |
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1341 |
lemma Max[simp]: assumes a: "finite S" "S \<noteq> {}"
|
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1342 |
shows "Max S \<in> S \<and> (\<forall>s \<in> S. s \<le> Max S)" (is "?P(Max S)") |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1343 |
proof (unfold Max_def, rule theI') |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1344 |
show "\<exists>!m. ?P m" |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1345 |
proof (rule ex_ex1I) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1346 |
show "\<exists>m. ?P m" using ex_Max[OF a] by blast |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1347 |
next |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1348 |
fix m1 m2 assume "?P m1" "?P m2" |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1349 |
thus "m1 = m2" by (blast dest:order_antisym) |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1350 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1351 |
qed |
|
44bdc150211b
Added Mi and Max on sets, hid Min and Pls on numerals.
nipkow
parents:
13421
diff
changeset
|
1352 |
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1353 |
subsection {* Theorems about @{text "choose"} *}
|
| 12396 | 1354 |
|
1355 |
text {*
|
|
1356 |
\medskip Basic theorem about @{text "choose"}. By Florian
|
|
1357 |
Kammüller, tidied by LCP. |
|
1358 |
*} |
|
1359 |
||
1360 |
lemma card_s_0_eq_empty: |
|
1361 |
"finite A ==> card {B. B \<subseteq> A & card B = 0} = 1"
|
|
1362 |
apply (simp cong add: conj_cong add: finite_subset [THEN card_0_eq]) |
|
1363 |
apply (simp cong add: rev_conj_cong) |
|
1364 |
done |
|
1365 |
||
1366 |
lemma choose_deconstruct: "finite M ==> x \<notin> M |
|
1367 |
==> {s. s <= insert x M & card(s) = Suc k}
|
|
1368 |
= {s. s <= M & card(s) = Suc k} Un
|
|
1369 |
{s. EX t. t <= M & card(t) = k & s = insert x t}"
|
|
1370 |
apply safe |
|
1371 |
apply (auto intro: finite_subset [THEN card_insert_disjoint]) |
|
1372 |
apply (drule_tac x = "xa - {x}" in spec)
|
|
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1373 |
apply (subgoal_tac "x \<notin> xa", auto) |
| 12396 | 1374 |
apply (erule rev_mp, subst card_Diff_singleton) |
1375 |
apply (auto intro: finite_subset) |
|
1376 |
done |
|
1377 |
||
1378 |
lemma card_inj_on_le: |
|
| 13595 | 1379 |
"[|inj_on f A; f ` A \<subseteq> B; finite A; finite B |] ==> card A <= card B" |
| 12396 | 1380 |
by (auto intro: card_mono simp add: card_image [symmetric]) |
1381 |
||
|
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1382 |
lemma card_bij_eq: |
|
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14331
diff
changeset
|
1383 |
"[|inj_on f A; f ` A \<subseteq> B; inj_on g B; g ` B \<subseteq> A; |
| 13595 | 1384 |
finite A; finite B |] ==> card A = card B" |
| 12396 | 1385 |
by (auto intro: le_anti_sym card_inj_on_le) |
1386 |
||
| 13595 | 1387 |
text{*There are as many subsets of @{term A} having cardinality @{term k}
|
1388 |
as there are sets obtained from the former by inserting a fixed element |
|
1389 |
@{term x} into each.*}
|
|
1390 |
lemma constr_bij: |
|
1391 |
"[|finite A; x \<notin> A|] ==> |
|
1392 |
card {B. EX C. C <= A & card(C) = k & B = insert x C} =
|
|
| 12396 | 1393 |
card {B. B <= A & card(B) = k}"
|
1394 |
apply (rule_tac f = "%s. s - {x}" and g = "insert x" in card_bij_eq)
|
|
| 13595 | 1395 |
apply (auto elim!: equalityE simp add: inj_on_def) |
1396 |
apply (subst Diff_insert0, auto) |
|
1397 |
txt {* finiteness of the two sets *}
|
|
1398 |
apply (rule_tac [2] B = "Pow (A)" in finite_subset) |
|
1399 |
apply (rule_tac B = "Pow (insert x A)" in finite_subset) |
|
1400 |
apply fast+ |
|
| 12396 | 1401 |
done |
1402 |
||
1403 |
text {*
|
|
1404 |
Main theorem: combinatorial statement about number of subsets of a set. |
|
1405 |
*} |
|
1406 |
||
1407 |
lemma n_sub_lemma: |
|
1408 |
"!!A. finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
|
|
1409 |
apply (induct k) |
|
| 14208 | 1410 |
apply (simp add: card_s_0_eq_empty, atomize) |
| 12396 | 1411 |
apply (rotate_tac -1, erule finite_induct) |
| 13421 | 1412 |
apply (simp_all (no_asm_simp) cong add: conj_cong |
1413 |
add: card_s_0_eq_empty choose_deconstruct) |
|
| 12396 | 1414 |
apply (subst card_Un_disjoint) |
1415 |
prefer 4 apply (force simp add: constr_bij) |
|
1416 |
prefer 3 apply force |
|
1417 |
prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2] |
|
1418 |
finite_subset [of _ "Pow (insert x F)", standard]) |
|
1419 |
apply (blast intro: finite_Pow_iff [THEN iffD2, THEN [2] finite_subset]) |
|
1420 |
done |
|
1421 |
||
| 13421 | 1422 |
theorem n_subsets: |
1423 |
"finite A ==> card {B. B <= A & card B = k} = (card A choose k)"
|
|
| 12396 | 1424 |
by (simp add: n_sub_lemma) |
1425 |
||
|
14443
75910c7557c5
generic theorems about exponentials; general tidying up
paulson
parents:
14430
diff
changeset
|
1426 |
|
| 12396 | 1427 |
end |