author | nipkow |
Mon, 22 Apr 1996 15:42:20 +0200 | |
changeset 1669 | e56cdf711729 |
parent 1660 | 8cb42cd97579 |
child 1748 | 88650ba93c10 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: HOL/equalities |
923 | 2 |
ID: $Id$ |
1465 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Equalities involving union, intersection, inclusion, etc. |
|
7 |
*) |
|
8 |
||
9 |
writeln"File HOL/equalities"; |
|
10 |
||
11 |
val eq_cs = set_cs addSIs [equalityI]; |
|
12 |
||
1548 | 13 |
section "{}"; |
14 |
||
1531 | 15 |
goal Set.thy "{x.False} = {}"; |
1553 | 16 |
by (fast_tac eq_cs 1); |
1531 | 17 |
qed "Collect_False_empty"; |
18 |
Addsimps [Collect_False_empty]; |
|
19 |
||
20 |
goal Set.thy "(A <= {}) = (A = {})"; |
|
1553 | 21 |
by (fast_tac eq_cs 1); |
1531 | 22 |
qed "subset_empty"; |
23 |
Addsimps [subset_empty]; |
|
24 |
||
1548 | 25 |
section ":"; |
923 | 26 |
|
27 |
goal Set.thy "x ~: {}"; |
|
1553 | 28 |
by (fast_tac set_cs 1); |
923 | 29 |
qed "in_empty"; |
1531 | 30 |
Addsimps[in_empty]; |
923 | 31 |
|
32 |
goal Set.thy "x : insert y A = (x=y | x:A)"; |
|
1553 | 33 |
by (fast_tac set_cs 1); |
923 | 34 |
qed "in_insert"; |
1531 | 35 |
Addsimps[in_insert]; |
923 | 36 |
|
1548 | 37 |
section "insert"; |
923 | 38 |
|
1531 | 39 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a {}*) |
40 |
goal Set.thy "insert a A = {a} Un A"; |
|
1553 | 41 |
by (fast_tac eq_cs 1); |
1531 | 42 |
qed "insert_is_Un"; |
43 |
||
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
44 |
goal Set.thy "insert a A ~= {}"; |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
45 |
by (fast_tac (set_cs addEs [equalityCE]) 1); |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
46 |
qed"insert_not_empty"; |
1531 | 47 |
Addsimps[insert_not_empty]; |
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
48 |
|
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
49 |
bind_thm("empty_not_insert",insert_not_empty RS not_sym); |
1531 | 50 |
Addsimps[empty_not_insert]; |
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
51 |
|
923 | 52 |
goal Set.thy "!!a. a:A ==> insert a A = A"; |
53 |
by (fast_tac eq_cs 1); |
|
54 |
qed "insert_absorb"; |
|
55 |
||
1531 | 56 |
goal Set.thy "insert x (insert x A) = insert x A"; |
1553 | 57 |
by (fast_tac eq_cs 1); |
1531 | 58 |
qed "insert_absorb2"; |
59 |
Addsimps [insert_absorb2]; |
|
60 |
||
923 | 61 |
goal Set.thy "(insert x A <= B) = (x:B & A <= B)"; |
62 |
by (fast_tac set_cs 1); |
|
63 |
qed "insert_subset"; |
|
1531 | 64 |
Addsimps[insert_subset]; |
65 |
||
66 |
(* use new B rather than (A-{a}) to avoid infinite unfolding *) |
|
67 |
goal Set.thy "!!a. a:A ==> ? B. A = insert a B & a ~: B"; |
|
1553 | 68 |
by (res_inst_tac [("x","A-{a}")] exI 1); |
69 |
by (fast_tac eq_cs 1); |
|
1531 | 70 |
qed "mk_disjoint_insert"; |
923 | 71 |
|
1660 | 72 |
section "``"; |
923 | 73 |
|
74 |
goal Set.thy "f``{} = {}"; |
|
75 |
by (fast_tac eq_cs 1); |
|
76 |
qed "image_empty"; |
|
1531 | 77 |
Addsimps[image_empty]; |
923 | 78 |
|
79 |
goal Set.thy "f``insert a B = insert (f a) (f``B)"; |
|
80 |
by (fast_tac eq_cs 1); |
|
81 |
qed "image_insert"; |
|
1531 | 82 |
Addsimps[image_insert]; |
923 | 83 |
|
1660 | 84 |
qed_goal "ball_image" Set.thy "(!y:F``S. P y) = (!x:S. P (F x))" |
85 |
(fn _ => [fast_tac set_cs 1]); |
|
86 |
||
87 |
section "range"; |
|
88 |
||
89 |
qed_goal "ball_range" Set.thy "(!y:range f. P y) = (!x. P (f x))" |
|
90 |
(fn _ => [fast_tac set_cs 1]); |
|
91 |
||
92 |
qed_goalw "image_range" Set.thy [image_def, range_def] |
|
93 |
"f``range g = range (%x. f (g x))" (fn _ => [ |
|
94 |
rtac Collect_cong 1, |
|
95 |
fast_tac set_cs 1]); |
|
96 |
||
1548 | 97 |
section "Int"; |
923 | 98 |
|
99 |
goal Set.thy "A Int A = A"; |
|
100 |
by (fast_tac eq_cs 1); |
|
101 |
qed "Int_absorb"; |
|
1531 | 102 |
Addsimps[Int_absorb]; |
923 | 103 |
|
104 |
goal Set.thy "A Int B = B Int A"; |
|
105 |
by (fast_tac eq_cs 1); |
|
106 |
qed "Int_commute"; |
|
107 |
||
108 |
goal Set.thy "(A Int B) Int C = A Int (B Int C)"; |
|
109 |
by (fast_tac eq_cs 1); |
|
110 |
qed "Int_assoc"; |
|
111 |
||
112 |
goal Set.thy "{} Int B = {}"; |
|
113 |
by (fast_tac eq_cs 1); |
|
114 |
qed "Int_empty_left"; |
|
1531 | 115 |
Addsimps[Int_empty_left]; |
923 | 116 |
|
117 |
goal Set.thy "A Int {} = {}"; |
|
118 |
by (fast_tac eq_cs 1); |
|
119 |
qed "Int_empty_right"; |
|
1531 | 120 |
Addsimps[Int_empty_right]; |
121 |
||
122 |
goal Set.thy "UNIV Int B = B"; |
|
123 |
by (fast_tac eq_cs 1); |
|
124 |
qed "Int_UNIV_left"; |
|
125 |
Addsimps[Int_UNIV_left]; |
|
126 |
||
127 |
goal Set.thy "A Int UNIV = A"; |
|
128 |
by (fast_tac eq_cs 1); |
|
129 |
qed "Int_UNIV_right"; |
|
130 |
Addsimps[Int_UNIV_right]; |
|
923 | 131 |
|
132 |
goal Set.thy "A Int (B Un C) = (A Int B) Un (A Int C)"; |
|
133 |
by (fast_tac eq_cs 1); |
|
134 |
qed "Int_Un_distrib"; |
|
135 |
||
1618 | 136 |
goal Set.thy "(B Un C) Int A = (B Int A) Un (C Int A)"; |
137 |
by (fast_tac eq_cs 1); |
|
138 |
qed "Int_Un_distrib2"; |
|
139 |
||
923 | 140 |
goal Set.thy "(A<=B) = (A Int B = A)"; |
141 |
by (fast_tac (eq_cs addSEs [equalityE]) 1); |
|
142 |
qed "subset_Int_eq"; |
|
143 |
||
1531 | 144 |
goal Set.thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)"; |
145 |
by (fast_tac (eq_cs addEs [equalityCE]) 1); |
|
146 |
qed "Int_UNIV"; |
|
147 |
Addsimps[Int_UNIV]; |
|
148 |
||
1548 | 149 |
section "Un"; |
923 | 150 |
|
151 |
goal Set.thy "A Un A = A"; |
|
152 |
by (fast_tac eq_cs 1); |
|
153 |
qed "Un_absorb"; |
|
1531 | 154 |
Addsimps[Un_absorb]; |
923 | 155 |
|
156 |
goal Set.thy "A Un B = B Un A"; |
|
157 |
by (fast_tac eq_cs 1); |
|
158 |
qed "Un_commute"; |
|
159 |
||
160 |
goal Set.thy "(A Un B) Un C = A Un (B Un C)"; |
|
161 |
by (fast_tac eq_cs 1); |
|
162 |
qed "Un_assoc"; |
|
163 |
||
164 |
goal Set.thy "{} Un B = B"; |
|
1553 | 165 |
by (fast_tac eq_cs 1); |
923 | 166 |
qed "Un_empty_left"; |
1531 | 167 |
Addsimps[Un_empty_left]; |
923 | 168 |
|
169 |
goal Set.thy "A Un {} = A"; |
|
1553 | 170 |
by (fast_tac eq_cs 1); |
923 | 171 |
qed "Un_empty_right"; |
1531 | 172 |
Addsimps[Un_empty_right]; |
173 |
||
174 |
goal Set.thy "UNIV Un B = UNIV"; |
|
1553 | 175 |
by (fast_tac eq_cs 1); |
1531 | 176 |
qed "Un_UNIV_left"; |
177 |
Addsimps[Un_UNIV_left]; |
|
178 |
||
179 |
goal Set.thy "A Un UNIV = UNIV"; |
|
1553 | 180 |
by (fast_tac eq_cs 1); |
1531 | 181 |
qed "Un_UNIV_right"; |
182 |
Addsimps[Un_UNIV_right]; |
|
923 | 183 |
|
184 |
goal Set.thy "insert a B Un C = insert a (B Un C)"; |
|
1553 | 185 |
by (fast_tac eq_cs 1); |
923 | 186 |
qed "Un_insert_left"; |
187 |
||
188 |
goal Set.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; |
|
189 |
by (fast_tac eq_cs 1); |
|
190 |
qed "Un_Int_distrib"; |
|
191 |
||
192 |
goal Set.thy |
|
193 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; |
|
194 |
by (fast_tac eq_cs 1); |
|
195 |
qed "Un_Int_crazy"; |
|
196 |
||
197 |
goal Set.thy "(A<=B) = (A Un B = B)"; |
|
198 |
by (fast_tac (eq_cs addSEs [equalityE]) 1); |
|
199 |
qed "subset_Un_eq"; |
|
200 |
||
201 |
goal Set.thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)"; |
|
202 |
by (fast_tac eq_cs 1); |
|
203 |
qed "subset_insert_iff"; |
|
204 |
||
205 |
goal Set.thy "(A Un B = {}) = (A = {} & B = {})"; |
|
206 |
by (fast_tac (eq_cs addEs [equalityCE]) 1); |
|
207 |
qed "Un_empty"; |
|
1531 | 208 |
Addsimps[Un_empty]; |
923 | 209 |
|
1548 | 210 |
section "Compl"; |
923 | 211 |
|
212 |
goal Set.thy "A Int Compl(A) = {}"; |
|
213 |
by (fast_tac eq_cs 1); |
|
214 |
qed "Compl_disjoint"; |
|
1531 | 215 |
Addsimps[Compl_disjoint]; |
923 | 216 |
|
1531 | 217 |
goal Set.thy "A Un Compl(A) = UNIV"; |
923 | 218 |
by (fast_tac eq_cs 1); |
219 |
qed "Compl_partition"; |
|
220 |
||
221 |
goal Set.thy "Compl(Compl(A)) = A"; |
|
222 |
by (fast_tac eq_cs 1); |
|
223 |
qed "double_complement"; |
|
1531 | 224 |
Addsimps[double_complement]; |
923 | 225 |
|
226 |
goal Set.thy "Compl(A Un B) = Compl(A) Int Compl(B)"; |
|
227 |
by (fast_tac eq_cs 1); |
|
228 |
qed "Compl_Un"; |
|
229 |
||
230 |
goal Set.thy "Compl(A Int B) = Compl(A) Un Compl(B)"; |
|
231 |
by (fast_tac eq_cs 1); |
|
232 |
qed "Compl_Int"; |
|
233 |
||
234 |
goal Set.thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; |
|
235 |
by (fast_tac eq_cs 1); |
|
236 |
qed "Compl_UN"; |
|
237 |
||
238 |
goal Set.thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; |
|
239 |
by (fast_tac eq_cs 1); |
|
240 |
qed "Compl_INT"; |
|
241 |
||
242 |
(*Halmos, Naive Set Theory, page 16.*) |
|
243 |
||
244 |
goal Set.thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)"; |
|
245 |
by (fast_tac (eq_cs addSEs [equalityE]) 1); |
|
246 |
qed "Un_Int_assoc_eq"; |
|
247 |
||
248 |
||
1548 | 249 |
section "Union"; |
923 | 250 |
|
251 |
goal Set.thy "Union({}) = {}"; |
|
252 |
by (fast_tac eq_cs 1); |
|
253 |
qed "Union_empty"; |
|
1531 | 254 |
Addsimps[Union_empty]; |
255 |
||
256 |
goal Set.thy "Union(UNIV) = UNIV"; |
|
257 |
by (fast_tac eq_cs 1); |
|
258 |
qed "Union_UNIV"; |
|
259 |
Addsimps[Union_UNIV]; |
|
923 | 260 |
|
261 |
goal Set.thy "Union(insert a B) = a Un Union(B)"; |
|
262 |
by (fast_tac eq_cs 1); |
|
263 |
qed "Union_insert"; |
|
1531 | 264 |
Addsimps[Union_insert]; |
923 | 265 |
|
266 |
goal Set.thy "Union(A Un B) = Union(A) Un Union(B)"; |
|
267 |
by (fast_tac eq_cs 1); |
|
268 |
qed "Union_Un_distrib"; |
|
1531 | 269 |
Addsimps[Union_Un_distrib]; |
923 | 270 |
|
271 |
goal Set.thy "Union(A Int B) <= Union(A) Int Union(B)"; |
|
272 |
by (fast_tac set_cs 1); |
|
273 |
qed "Union_Int_subset"; |
|
274 |
||
275 |
val prems = goal Set.thy |
|
276 |
"(Union(C) Int A = {}) = (! B:C. B Int A = {})"; |
|
277 |
by (fast_tac (eq_cs addSEs [equalityE]) 1); |
|
278 |
qed "Union_disjoint"; |
|
279 |
||
1548 | 280 |
section "Inter"; |
281 |
||
1531 | 282 |
goal Set.thy "Inter({}) = UNIV"; |
283 |
by (fast_tac eq_cs 1); |
|
284 |
qed "Inter_empty"; |
|
285 |
Addsimps[Inter_empty]; |
|
286 |
||
287 |
goal Set.thy "Inter(UNIV) = {}"; |
|
288 |
by (fast_tac eq_cs 1); |
|
289 |
qed "Inter_UNIV"; |
|
290 |
Addsimps[Inter_UNIV]; |
|
291 |
||
292 |
goal Set.thy "Inter(insert a B) = a Int Inter(B)"; |
|
293 |
by (fast_tac eq_cs 1); |
|
294 |
qed "Inter_insert"; |
|
295 |
Addsimps[Inter_insert]; |
|
296 |
||
1564
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents:
1553
diff
changeset
|
297 |
goal Set.thy "Inter(A) Un Inter(B) <= Inter(A Int B)"; |
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents:
1553
diff
changeset
|
298 |
by (fast_tac set_cs 1); |
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents:
1553
diff
changeset
|
299 |
qed "Inter_Un_subset"; |
1531 | 300 |
|
923 | 301 |
goal Set.thy "Inter(A Un B) = Inter(A) Int Inter(B)"; |
302 |
by (best_tac eq_cs 1); |
|
303 |
qed "Inter_Un_distrib"; |
|
304 |
||
1548 | 305 |
section "UN and INT"; |
923 | 306 |
|
307 |
(*Basic identities*) |
|
308 |
||
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
309 |
goal Set.thy "(UN x:{}. B x) = {}"; |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
310 |
by (fast_tac eq_cs 1); |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
311 |
qed "UN_empty"; |
1531 | 312 |
Addsimps[UN_empty]; |
313 |
||
314 |
goal Set.thy "(UN x:UNIV. B x) = (UN x. B x)"; |
|
315 |
by (fast_tac eq_cs 1); |
|
316 |
qed "UN_UNIV"; |
|
317 |
Addsimps[UN_UNIV]; |
|
318 |
||
319 |
goal Set.thy "(INT x:{}. B x) = UNIV"; |
|
320 |
by (fast_tac eq_cs 1); |
|
321 |
qed "INT_empty"; |
|
322 |
Addsimps[INT_empty]; |
|
323 |
||
324 |
goal Set.thy "(INT x:UNIV. B x) = (INT x. B x)"; |
|
325 |
by (fast_tac eq_cs 1); |
|
326 |
qed "INT_UNIV"; |
|
327 |
Addsimps[INT_UNIV]; |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
328 |
|
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
329 |
goal Set.thy "(UN x:insert a A. B x) = B a Un UNION A B"; |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
330 |
by (fast_tac eq_cs 1); |
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
331 |
qed "UN_insert"; |
1531 | 332 |
Addsimps[UN_insert]; |
333 |
||
334 |
goal Set.thy "(INT x:insert a A. B x) = B a Int INTER A B"; |
|
335 |
by (fast_tac eq_cs 1); |
|
336 |
qed "INT_insert"; |
|
337 |
Addsimps[INT_insert]; |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
338 |
|
923 | 339 |
goal Set.thy "Union(range(f)) = (UN x.f(x))"; |
340 |
by (fast_tac eq_cs 1); |
|
341 |
qed "Union_range_eq"; |
|
342 |
||
343 |
goal Set.thy "Inter(range(f)) = (INT x.f(x))"; |
|
344 |
by (fast_tac eq_cs 1); |
|
345 |
qed "Inter_range_eq"; |
|
346 |
||
347 |
goal Set.thy "Union(B``A) = (UN x:A. B(x))"; |
|
348 |
by (fast_tac eq_cs 1); |
|
349 |
qed "Union_image_eq"; |
|
350 |
||
351 |
goal Set.thy "Inter(B``A) = (INT x:A. B(x))"; |
|
352 |
by (fast_tac eq_cs 1); |
|
353 |
qed "Inter_image_eq"; |
|
354 |
||
355 |
goal Set.thy "!!A. a: A ==> (UN y:A. c) = c"; |
|
356 |
by (fast_tac eq_cs 1); |
|
357 |
qed "UN_constant"; |
|
358 |
||
359 |
goal Set.thy "!!A. a: A ==> (INT y:A. c) = c"; |
|
360 |
by (fast_tac eq_cs 1); |
|
361 |
qed "INT_constant"; |
|
362 |
||
363 |
goal Set.thy "(UN x.B) = B"; |
|
364 |
by (fast_tac eq_cs 1); |
|
365 |
qed "UN1_constant"; |
|
1531 | 366 |
Addsimps[UN1_constant]; |
923 | 367 |
|
368 |
goal Set.thy "(INT x.B) = B"; |
|
369 |
by (fast_tac eq_cs 1); |
|
370 |
qed "INT1_constant"; |
|
1531 | 371 |
Addsimps[INT1_constant]; |
923 | 372 |
|
373 |
goal Set.thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})"; |
|
374 |
by (fast_tac eq_cs 1); |
|
375 |
qed "UN_eq"; |
|
376 |
||
377 |
(*Look: it has an EXISTENTIAL quantifier*) |
|
378 |
goal Set.thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})"; |
|
379 |
by (fast_tac eq_cs 1); |
|
380 |
qed "INT_eq"; |
|
381 |
||
382 |
(*Distributive laws...*) |
|
383 |
||
384 |
goal Set.thy "A Int Union(B) = (UN C:B. A Int C)"; |
|
385 |
by (fast_tac eq_cs 1); |
|
386 |
qed "Int_Union"; |
|
387 |
||
388 |
(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: |
|
389 |
Union of a family of unions **) |
|
390 |
goal Set.thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; |
|
391 |
by (fast_tac eq_cs 1); |
|
392 |
qed "Un_Union_image"; |
|
393 |
||
394 |
(*Equivalent version*) |
|
395 |
goal Set.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; |
|
396 |
by (fast_tac eq_cs 1); |
|
397 |
qed "UN_Un_distrib"; |
|
398 |
||
399 |
goal Set.thy "A Un Inter(B) = (INT C:B. A Un C)"; |
|
400 |
by (fast_tac eq_cs 1); |
|
401 |
qed "Un_Inter"; |
|
402 |
||
403 |
goal Set.thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; |
|
404 |
by (best_tac eq_cs 1); |
|
405 |
qed "Int_Inter_image"; |
|
406 |
||
407 |
(*Equivalent version*) |
|
408 |
goal Set.thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; |
|
409 |
by (fast_tac eq_cs 1); |
|
410 |
qed "INT_Int_distrib"; |
|
411 |
||
412 |
(*Halmos, Naive Set Theory, page 35.*) |
|
413 |
goal Set.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; |
|
414 |
by (fast_tac eq_cs 1); |
|
415 |
qed "Int_UN_distrib"; |
|
416 |
||
417 |
goal Set.thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; |
|
418 |
by (fast_tac eq_cs 1); |
|
419 |
qed "Un_INT_distrib"; |
|
420 |
||
421 |
goal Set.thy |
|
422 |
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; |
|
423 |
by (fast_tac eq_cs 1); |
|
424 |
qed "Int_UN_distrib2"; |
|
425 |
||
426 |
goal Set.thy |
|
427 |
"(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; |
|
428 |
by (fast_tac eq_cs 1); |
|
429 |
qed "Un_INT_distrib2"; |
|
430 |
||
1548 | 431 |
section "-"; |
923 | 432 |
|
433 |
goal Set.thy "A-A = {}"; |
|
434 |
by (fast_tac eq_cs 1); |
|
435 |
qed "Diff_cancel"; |
|
1531 | 436 |
Addsimps[Diff_cancel]; |
923 | 437 |
|
438 |
goal Set.thy "{}-A = {}"; |
|
439 |
by (fast_tac eq_cs 1); |
|
440 |
qed "empty_Diff"; |
|
1531 | 441 |
Addsimps[empty_Diff]; |
923 | 442 |
|
443 |
goal Set.thy "A-{} = A"; |
|
444 |
by (fast_tac eq_cs 1); |
|
445 |
qed "Diff_empty"; |
|
1531 | 446 |
Addsimps[Diff_empty]; |
447 |
||
448 |
goal Set.thy "A-UNIV = {}"; |
|
449 |
by (fast_tac eq_cs 1); |
|
450 |
qed "Diff_UNIV"; |
|
451 |
Addsimps[Diff_UNIV]; |
|
452 |
||
453 |
goal Set.thy "!!x. x~:A ==> A - insert x B = A-B"; |
|
1553 | 454 |
by (fast_tac eq_cs 1); |
1531 | 455 |
qed "Diff_insert0"; |
456 |
Addsimps [Diff_insert0]; |
|
923 | 457 |
|
458 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) |
|
459 |
goal Set.thy "A - insert a B = A - B - {a}"; |
|
460 |
by (fast_tac eq_cs 1); |
|
461 |
qed "Diff_insert"; |
|
462 |
||
463 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*) |
|
464 |
goal Set.thy "A - insert a B = A - {a} - B"; |
|
465 |
by (fast_tac eq_cs 1); |
|
466 |
qed "Diff_insert2"; |
|
467 |
||
1531 | 468 |
goal Set.thy "insert x A - B = (if x:B then A-B else insert x (A-B))"; |
1553 | 469 |
by (simp_tac (!simpset setloop split_tac[expand_if]) 1); |
470 |
by (fast_tac eq_cs 1); |
|
1531 | 471 |
qed "insert_Diff_if"; |
472 |
||
473 |
goal Set.thy "!!x. x:B ==> insert x A - B = A-B"; |
|
1553 | 474 |
by (fast_tac eq_cs 1); |
1531 | 475 |
qed "insert_Diff1"; |
476 |
Addsimps [insert_Diff1]; |
|
477 |
||
923 | 478 |
val prems = goal Set.thy "a:A ==> insert a (A-{a}) = A"; |
479 |
by (fast_tac (eq_cs addSIs prems) 1); |
|
480 |
qed "insert_Diff"; |
|
481 |
||
482 |
goal Set.thy "A Int (B-A) = {}"; |
|
483 |
by (fast_tac eq_cs 1); |
|
484 |
qed "Diff_disjoint"; |
|
1531 | 485 |
Addsimps[Diff_disjoint]; |
923 | 486 |
|
487 |
goal Set.thy "!!A. A<=B ==> A Un (B-A) = B"; |
|
488 |
by (fast_tac eq_cs 1); |
|
489 |
qed "Diff_partition"; |
|
490 |
||
491 |
goal Set.thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; |
|
492 |
by (fast_tac eq_cs 1); |
|
493 |
qed "double_diff"; |
|
494 |
||
495 |
goal Set.thy "A - (B Un C) = (A-B) Int (A-C)"; |
|
496 |
by (fast_tac eq_cs 1); |
|
497 |
qed "Diff_Un"; |
|
498 |
||
499 |
goal Set.thy "A - (B Int C) = (A-B) Un (A-C)"; |
|
500 |
by (fast_tac eq_cs 1); |
|
501 |
qed "Diff_Int"; |
|
502 |
||
1531 | 503 |
Addsimps[subset_UNIV, empty_subsetI, subset_refl]; |