author | wenzelm |
Tue, 11 Dec 2001 16:00:26 +0100 | |
changeset 12466 | 5f4182667032 |
parent 12199 | 8213fd95acb5 |
child 12891 | 92af5c3a10fb |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/equalities |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
Set Theory examples: Union, Intersection, Inclusion, etc. |
|
7 |
(Thanks also to Philippe de Groote.) |
|
8 |
*) |
|
9 |
||
10 |
(** Finite Sets **) |
|
11 |
||
520 | 12 |
(* cons_def refers to Upair; reversing the equality LOOPS in rewriting!*) |
5321 | 13 |
Goal "{a} Un B = cons(a,B)"; |
2877 | 14 |
by (Blast_tac 1); |
760 | 15 |
qed "cons_eq"; |
520 | 16 |
|
5321 | 17 |
Goal "cons(a, cons(b, C)) = cons(b, cons(a, C))"; |
2877 | 18 |
by (Blast_tac 1); |
760 | 19 |
qed "cons_commute"; |
0 | 20 |
|
5321 | 21 |
Goal "a: B ==> cons(a,B) = B"; |
2877 | 22 |
by (Blast_tac 1); |
760 | 23 |
qed "cons_absorb"; |
0 | 24 |
|
5321 | 25 |
Goal "a: B ==> cons(a, B-{a}) = B"; |
2877 | 26 |
by (Blast_tac 1); |
760 | 27 |
qed "cons_Diff"; |
0 | 28 |
|
5321 | 29 |
Goal "[| a: C; ALL y:C. y=b |] ==> C = {b}"; |
2877 | 30 |
by (Blast_tac 1); |
760 | 31 |
qed "equal_singleton_lemma"; |
9907 | 32 |
bind_thm ("equal_singleton", ballI RSN (2,equal_singleton_lemma)); |
0 | 33 |
|
34 |
||
35 |
(** Binary Intersection **) |
|
36 |
||
37 |
(*NOT an equality, but it seems to belong here...*) |
|
5321 | 38 |
Goal "cons(a,B) Int C <= cons(a, B Int C)"; |
2877 | 39 |
by (Blast_tac 1); |
760 | 40 |
qed "Int_cons"; |
0 | 41 |
|
5321 | 42 |
Goal "A Int A = A"; |
2877 | 43 |
by (Blast_tac 1); |
760 | 44 |
qed "Int_absorb"; |
12199 | 45 |
Addsimps [Int_absorb]; |
46 |
||
47 |
Goal "A Int (A Int B) = A Int B"; |
|
48 |
by (Blast_tac 1); |
|
49 |
qed "Int_left_absorb"; |
|
0 | 50 |
|
5321 | 51 |
Goal "A Int B = B Int A"; |
2877 | 52 |
by (Blast_tac 1); |
760 | 53 |
qed "Int_commute"; |
0 | 54 |
|
12199 | 55 |
Goal "A Int (B Int C) = B Int (A Int C)"; |
56 |
by (Blast_tac 1); |
|
57 |
qed "Int_left_commute"; |
|
58 |
||
5321 | 59 |
Goal "(A Int B) Int C = A Int (B Int C)"; |
2877 | 60 |
by (Blast_tac 1); |
760 | 61 |
qed "Int_assoc"; |
0 | 62 |
|
12199 | 63 |
(*Intersection is an AC-operator*) |
64 |
bind_thms ("Int_ac", |
|
65 |
[Int_assoc, Int_left_absorb, Int_commute, Int_left_commute]); |
|
66 |
||
67 |
Goal "A Int (B Un C) = (A Int B) Un (A Int C)"; |
|
2877 | 68 |
by (Blast_tac 1); |
760 | 69 |
qed "Int_Un_distrib"; |
0 | 70 |
|
12199 | 71 |
Goal "(B Un C) Int A = (B Int A) Un (C Int A)"; |
72 |
by (Blast_tac 1); |
|
73 |
qed "Int_Un_distrib2"; |
|
74 |
||
5321 | 75 |
Goal "A<=B <-> A Int B = A"; |
4091 | 76 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 77 |
qed "subset_Int_iff"; |
0 | 78 |
|
5321 | 79 |
Goal "A<=B <-> B Int A = A"; |
4091 | 80 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 81 |
qed "subset_Int_iff2"; |
435 | 82 |
|
5321 | 83 |
Goal "C<=A ==> (A-B) Int C = C-B"; |
2877 | 84 |
by (Blast_tac 1); |
1035 | 85 |
qed "Int_Diff_eq"; |
86 |
||
0 | 87 |
(** Binary Union **) |
88 |
||
5321 | 89 |
Goal "cons(a,B) Un C = cons(a, B Un C)"; |
2877 | 90 |
by (Blast_tac 1); |
760 | 91 |
qed "Un_cons"; |
0 | 92 |
|
5321 | 93 |
Goal "A Un A = A"; |
2877 | 94 |
by (Blast_tac 1); |
760 | 95 |
qed "Un_absorb"; |
12199 | 96 |
Addsimps [Un_absorb]; |
97 |
||
98 |
Goal "A Un (A Un B) = A Un B"; |
|
99 |
by (Blast_tac 1); |
|
100 |
qed "Un_left_absorb"; |
|
0 | 101 |
|
5321 | 102 |
Goal "A Un B = B Un A"; |
2877 | 103 |
by (Blast_tac 1); |
760 | 104 |
qed "Un_commute"; |
0 | 105 |
|
12199 | 106 |
Goal "A Un (B Un C) = B Un (A Un C)"; |
107 |
by (Blast_tac 1); |
|
108 |
qed "Un_left_commute"; |
|
109 |
||
5321 | 110 |
Goal "(A Un B) Un C = A Un (B Un C)"; |
2877 | 111 |
by (Blast_tac 1); |
760 | 112 |
qed "Un_assoc"; |
0 | 113 |
|
12199 | 114 |
(*Union is an AC-operator*) |
115 |
bind_thms ("Un_ac", [Un_assoc, Un_left_absorb, Un_commute, Un_left_commute]); |
|
116 |
||
5321 | 117 |
Goal "(A Int B) Un C = (A Un C) Int (B Un C)"; |
2877 | 118 |
by (Blast_tac 1); |
760 | 119 |
qed "Un_Int_distrib"; |
0 | 120 |
|
5321 | 121 |
Goal "A<=B <-> A Un B = B"; |
4091 | 122 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 123 |
qed "subset_Un_iff"; |
0 | 124 |
|
5321 | 125 |
Goal "A<=B <-> B Un A = B"; |
4091 | 126 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 127 |
qed "subset_Un_iff2"; |
435 | 128 |
|
12199 | 129 |
Goal "(A Un B = 0) <-> (A = 0 & B = 0)"; |
130 |
by (Blast_tac 1); |
|
131 |
qed "Un_empty"; |
|
132 |
AddIffs[Un_empty]; |
|
133 |
||
134 |
Goal "A Un B = Union({A, B})"; |
|
135 |
by (Blast_tac 1); |
|
136 |
qed "Un_eq_Union"; |
|
137 |
||
0 | 138 |
(** Simple properties of Diff -- set difference **) |
139 |
||
11695 | 140 |
Goal "A - A = 0"; |
2877 | 141 |
by (Blast_tac 1); |
760 | 142 |
qed "Diff_cancel"; |
0 | 143 |
|
12199 | 144 |
Goal "A Int B = 0 ==> A - B = A"; |
145 |
by (Blast_tac 1); |
|
146 |
qed "Diff_triv"; |
|
147 |
||
11695 | 148 |
Goal "0 - A = 0"; |
2877 | 149 |
by (Blast_tac 1); |
760 | 150 |
qed "empty_Diff"; |
12199 | 151 |
Addsimps[empty_Diff]; |
0 | 152 |
|
11695 | 153 |
Goal "A - 0 = A"; |
2877 | 154 |
by (Blast_tac 1); |
760 | 155 |
qed "Diff_0"; |
12199 | 156 |
Addsimps[Diff_0]; |
0 | 157 |
|
11695 | 158 |
Goal "A - B = 0 <-> A <= B"; |
4091 | 159 |
by (blast_tac (claset() addEs [equalityE]) 1); |
787 | 160 |
qed "Diff_eq_0_iff"; |
161 |
||
0 | 162 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
5321 | 163 |
Goal "A - cons(a,B) = A - B - {a}"; |
2877 | 164 |
by (Blast_tac 1); |
760 | 165 |
qed "Diff_cons"; |
0 | 166 |
|
167 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
|
5321 | 168 |
Goal "A - cons(a,B) = A - {a} - B"; |
2877 | 169 |
by (Blast_tac 1); |
760 | 170 |
qed "Diff_cons2"; |
0 | 171 |
|
5321 | 172 |
Goal "A Int (B-A) = 0"; |
2877 | 173 |
by (Blast_tac 1); |
760 | 174 |
qed "Diff_disjoint"; |
0 | 175 |
|
5321 | 176 |
Goal "A<=B ==> A Un (B-A) = B"; |
2877 | 177 |
by (Blast_tac 1); |
760 | 178 |
qed "Diff_partition"; |
0 | 179 |
|
5321 | 180 |
Goal "A <= B Un (A - B)"; |
2877 | 181 |
by (Blast_tac 1); |
1611 | 182 |
qed "subset_Un_Diff"; |
183 |
||
5321 | 184 |
Goal "[| A<=B; B<=C |] ==> B-(C-A) = A"; |
2877 | 185 |
by (Blast_tac 1); |
760 | 186 |
qed "double_complement"; |
0 | 187 |
|
5321 | 188 |
Goal "(A Un B) - (B-A) = A"; |
2877 | 189 |
by (Blast_tac 1); |
760 | 190 |
qed "double_complement_Un"; |
268 | 191 |
|
5321 | 192 |
Goal |
0 | 193 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; |
2877 | 194 |
by (Blast_tac 1); |
760 | 195 |
qed "Un_Int_crazy"; |
0 | 196 |
|
5321 | 197 |
Goal "A - (B Un C) = (A-B) Int (A-C)"; |
2877 | 198 |
by (Blast_tac 1); |
760 | 199 |
qed "Diff_Un"; |
0 | 200 |
|
5321 | 201 |
Goal "A - (B Int C) = (A-B) Un (A-C)"; |
2877 | 202 |
by (Blast_tac 1); |
760 | 203 |
qed "Diff_Int"; |
0 | 204 |
|
12199 | 205 |
Goal "(A Un B) - C = (A - C) Un (B - C)"; |
206 |
by (Blast_tac 1); |
|
207 |
qed "Un_Diff"; |
|
208 |
||
209 |
Goal "(A Int B) - C = A Int (B - C)"; |
|
210 |
by (Blast_tac 1); |
|
211 |
qed "Int_Diff"; |
|
212 |
||
213 |
Goal "C Int (A-B) = (C Int A) - (C Int B)"; |
|
214 |
by (Blast_tac 1); |
|
215 |
qed "Diff_Int_distrib"; |
|
216 |
||
217 |
Goal "(A-B) Int C = (A Int C) - (B Int C)"; |
|
218 |
by (Blast_tac 1); |
|
219 |
qed "Diff_Int_distrib2"; |
|
220 |
||
0 | 221 |
(*Halmos, Naive Set Theory, page 16.*) |
5321 | 222 |
Goal "(A Int B) Un C = A Int (B Un C) <-> C<=A"; |
4091 | 223 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 224 |
qed "Un_Int_assoc_iff"; |
0 | 225 |
|
226 |
||
227 |
(** Big Union and Intersection **) |
|
228 |
||
5321 | 229 |
Goal "Union(cons(a,B)) = a Un Union(B)"; |
2877 | 230 |
by (Blast_tac 1); |
760 | 231 |
qed "Union_cons"; |
12199 | 232 |
Addsimps [Union_cons]; |
0 | 233 |
|
5321 | 234 |
Goal "Union(A Un B) = Union(A) Un Union(B)"; |
2877 | 235 |
by (Blast_tac 1); |
760 | 236 |
qed "Union_Un_distrib"; |
0 | 237 |
|
5321 | 238 |
Goal "Union(A Int B) <= Union(A) Int Union(B)"; |
2877 | 239 |
by (Blast_tac 1); |
760 | 240 |
qed "Union_Int_subset"; |
435 | 241 |
|
5321 | 242 |
Goal "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)"; |
4091 | 243 |
by (blast_tac (claset() addSEs [equalityE]) 1); |
760 | 244 |
qed "Union_disjoint"; |
0 | 245 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
246 |
Goal "Union(A) = 0 <-> (ALL B:A. B=0)"; |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
247 |
by (Blast_tac 1); |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
248 |
qed "Union_empty_iff"; |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
249 |
|
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
250 |
Goalw [Inter_def] "Inter(0) = 0"; |
2877 | 251 |
by (Blast_tac 1); |
1652 | 252 |
qed "Inter_0"; |
253 |
||
5321 | 254 |
Goal "[| z:A; z:B |] ==> Inter(A) Un Inter(B) <= Inter(A Int B)"; |
2877 | 255 |
by (Blast_tac 1); |
1568 | 256 |
qed "Inter_Un_subset"; |
257 |
||
0 | 258 |
(* A good challenge: Inter is ill-behaved on the empty set *) |
5321 | 259 |
Goal "[| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)"; |
2877 | 260 |
by (Blast_tac 1); |
760 | 261 |
qed "Inter_Un_distrib"; |
0 | 262 |
|
5321 | 263 |
Goal "Union({b}) = b"; |
2877 | 264 |
by (Blast_tac 1); |
760 | 265 |
qed "Union_singleton"; |
0 | 266 |
|
5321 | 267 |
Goal "Inter({b}) = b"; |
2877 | 268 |
by (Blast_tac 1); |
760 | 269 |
qed "Inter_singleton"; |
0 | 270 |
|
12199 | 271 |
Goal "Inter(cons(a,B)) = (if B=0 then a else a Int Inter(B))"; |
272 |
by (Simp_tac 1); |
|
273 |
by (Blast_tac 1); |
|
274 |
qed "Inter_cons"; |
|
275 |
Addsimps [Inter_cons]; |
|
276 |
||
0 | 277 |
(** Unions and Intersections of Families **) |
278 |
||
5321 | 279 |
Goal "Union(A) = (UN x:A. x)"; |
2877 | 280 |
by (Blast_tac 1); |
760 | 281 |
qed "Union_eq_UN"; |
0 | 282 |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
283 |
Goalw [Inter_def] "Inter(A) = (INT x:A. x)"; |
2877 | 284 |
by (Blast_tac 1); |
760 | 285 |
qed "Inter_eq_INT"; |
0 | 286 |
|
5321 | 287 |
Goal "(UN i:0. A(i)) = 0"; |
2877 | 288 |
by (Blast_tac 1); |
760 | 289 |
qed "UN_0"; |
12199 | 290 |
Addsimps [UN_0]; |
291 |
||
292 |
Goal "(UN x:A. {x}) = A"; |
|
293 |
by (Blast_tac 1); |
|
294 |
qed "UN_singleton"; |
|
295 |
||
296 |
Goal "(UN i: A Un B. C(i)) = (UN i: A. C(i)) Un (UN i:B. C(i))"; |
|
297 |
by (Blast_tac 1); |
|
298 |
qed "UN_Un"; |
|
299 |
||
300 |
Goal "(INT i:I Un J. A(i)) = (if I=0 then INT j:J. A(j) \ |
|
301 |
\ else if J=0 then INT i:I. A(i) \ |
|
302 |
\ else ((INT i:I. A(i)) Int (INT j:J. A(j))))"; |
|
303 |
by Auto_tac; |
|
304 |
by (blast_tac (claset() addSIs [equalityI]) 1); |
|
305 |
qed "INT_Un"; |
|
306 |
||
307 |
Goal "(UN x : (UN y:A. B(y)). C(x)) = (UN y:A. UN x: B(y). C(x))"; |
|
308 |
by (Blast_tac 1); |
|
309 |
qed "UN_UN_flatten"; |
|
517 | 310 |
|
0 | 311 |
(*Halmos, Naive Set Theory, page 35.*) |
5321 | 312 |
Goal "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; |
2877 | 313 |
by (Blast_tac 1); |
760 | 314 |
qed "Int_UN_distrib"; |
0 | 315 |
|
5321 | 316 |
Goal "i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; |
2877 | 317 |
by (Blast_tac 1); |
760 | 318 |
qed "Un_INT_distrib"; |
0 | 319 |
|
5321 | 320 |
Goal "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; |
2877 | 321 |
by (Blast_tac 1); |
760 | 322 |
qed "Int_UN_distrib2"; |
0 | 323 |
|
5321 | 324 |
Goal "[| i:I; j:J |] ==> \ |
325 |
\ (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; |
|
2877 | 326 |
by (Blast_tac 1); |
760 | 327 |
qed "Un_INT_distrib2"; |
0 | 328 |
|
5321 | 329 |
Goal "a: A ==> (UN y:A. c) = c"; |
2877 | 330 |
by (Blast_tac 1); |
760 | 331 |
qed "UN_constant"; |
0 | 332 |
|
5321 | 333 |
Goal "a: A ==> (INT y:A. c) = c"; |
2877 | 334 |
by (Blast_tac 1); |
760 | 335 |
qed "INT_constant"; |
0 | 336 |
|
5321 | 337 |
Goal "(UN y: RepFun(A,f). B(y)) = (UN x:A. B(f(x)))"; |
4242 | 338 |
by (Blast_tac 1); |
339 |
qed "UN_RepFun"; |
|
12199 | 340 |
Addsimps [UN_RepFun]; |
4242 | 341 |
|
12199 | 342 |
Goal "(INT x:RepFun(A,f). B(x)) = (INT a:A. B(f(a)))"; |
343 |
by (auto_tac (claset(), simpset() addsimps [Inter_def])); |
|
344 |
qed "INT_RepFun"; |
|
345 |
Addsimps [INT_RepFun]; |
|
346 |
||
347 |
Goal "0 ~: A ==> (INT x: Union(A). B(x)) = (INT y:A. INT x:y. B(x))"; |
|
348 |
by (asm_full_simp_tac (simpset() addsimps [Inter_def]) 1); |
|
349 |
by (subgoal_tac "ALL x:A. x~=0" 1); |
|
350 |
by (Blast_tac 2); |
|
351 |
by (rtac equalityI 1); |
|
352 |
by (Clarify_tac 1); |
|
353 |
by (blast_tac (claset() addIs []) 1); |
|
354 |
by (blast_tac (claset() addSDs [bspec]) 1); |
|
355 |
qed "INT_Union_eq"; |
|
356 |
||
357 |
Goal "(ALL x:A. B(x) ~= 0) \ |
|
358 |
\ ==> (INT z: (UN x:A. B(x)). C(z)) = (INT x:A. INT z: B(x). C(z))"; |
|
359 |
by (stac INT_Union_eq 1); |
|
4242 | 360 |
by (Blast_tac 1); |
12199 | 361 |
by (simp_tac (simpset() addsimps [Inter_def]) 1); |
362 |
qed "INT_UN_eq"; |
|
4242 | 363 |
|
364 |
||
0 | 365 |
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: |
366 |
Union of a family of unions **) |
|
367 |
||
5321 | 368 |
Goal "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; |
2877 | 369 |
by (Blast_tac 1); |
760 | 370 |
qed "UN_Un_distrib"; |
0 | 371 |
|
5321 | 372 |
Goal "i:I ==> (INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; |
2877 | 373 |
by (Blast_tac 1); |
760 | 374 |
qed "INT_Int_distrib"; |
0 | 375 |
|
5321 | 376 |
Goal "(UN z:I Int J. A(z)) <= (UN z:I. A(z)) Int (UN z:J. A(z))"; |
2877 | 377 |
by (Blast_tac 1); |
1784 | 378 |
qed "UN_Int_subset"; |
379 |
||
0 | 380 |
(** Devlin, page 12, exercise 5: Complements **) |
381 |
||
5321 | 382 |
Goal "i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))"; |
2877 | 383 |
by (Blast_tac 1); |
760 | 384 |
qed "Diff_UN"; |
0 | 385 |
|
5321 | 386 |
Goal "i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))"; |
2877 | 387 |
by (Blast_tac 1); |
760 | 388 |
qed "Diff_INT"; |
0 | 389 |
|
390 |
(** Unions and Intersections with General Sum **) |
|
391 |
||
1611 | 392 |
(*Not suitable for rewriting: LOOPS!*) |
5321 | 393 |
Goal "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)"; |
2877 | 394 |
by (Blast_tac 1); |
1611 | 395 |
qed "Sigma_cons1"; |
396 |
||
397 |
(*Not suitable for rewriting: LOOPS!*) |
|
5321 | 398 |
Goal "A * cons(b,B) = A*{b} Un A*B"; |
2877 | 399 |
by (Blast_tac 1); |
1611 | 400 |
qed "Sigma_cons2"; |
401 |
||
5321 | 402 |
Goal "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)"; |
2877 | 403 |
by (Blast_tac 1); |
1611 | 404 |
qed "Sigma_succ1"; |
405 |
||
5321 | 406 |
Goal "A * succ(B) = A*{B} Un A*B"; |
2877 | 407 |
by (Blast_tac 1); |
1611 | 408 |
qed "Sigma_succ2"; |
520 | 409 |
|
5321 | 410 |
Goal "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))"; |
2877 | 411 |
by (Blast_tac 1); |
760 | 412 |
qed "SUM_UN_distrib1"; |
182 | 413 |
|
5321 | 414 |
Goal "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))"; |
2877 | 415 |
by (Blast_tac 1); |
760 | 416 |
qed "SUM_UN_distrib2"; |
182 | 417 |
|
5321 | 418 |
Goal "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))"; |
2877 | 419 |
by (Blast_tac 1); |
760 | 420 |
qed "SUM_Un_distrib1"; |
0 | 421 |
|
5321 | 422 |
Goal "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))"; |
2877 | 423 |
by (Blast_tac 1); |
760 | 424 |
qed "SUM_Un_distrib2"; |
0 | 425 |
|
685 | 426 |
(*First-order version of the above, for rewriting*) |
5321 | 427 |
Goal "I * (A Un B) = I*A Un I*B"; |
1461 | 428 |
by (rtac SUM_Un_distrib2 1); |
760 | 429 |
qed "prod_Un_distrib2"; |
685 | 430 |
|
5321 | 431 |
Goal "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))"; |
2877 | 432 |
by (Blast_tac 1); |
760 | 433 |
qed "SUM_Int_distrib1"; |
0 | 434 |
|
5321 | 435 |
Goal "(SUM i:I. A(i) Int B(i)) = (SUM i:I. A(i)) Int (SUM i:I. B(i))"; |
2877 | 436 |
by (Blast_tac 1); |
760 | 437 |
qed "SUM_Int_distrib2"; |
0 | 438 |
|
685 | 439 |
(*First-order version of the above, for rewriting*) |
5321 | 440 |
Goal "I * (A Int B) = I*A Int I*B"; |
1461 | 441 |
by (rtac SUM_Int_distrib2 1); |
760 | 442 |
qed "prod_Int_distrib2"; |
685 | 443 |
|
192 | 444 |
(*Cf Aczel, Non-Well-Founded Sets, page 115*) |
5321 | 445 |
Goal "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))"; |
2877 | 446 |
by (Blast_tac 1); |
760 | 447 |
qed "SUM_eq_UN"; |
192 | 448 |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
449 |
(** Domain **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
450 |
|
5321 | 451 |
Goal "b:B ==> domain(A*B) = A"; |
452 |
by (Blast_tac 1); |
|
453 |
qed "domain_of_prod"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
454 |
|
9180 | 455 |
Goal "domain(0) = 0"; |
456 |
by (Blast_tac 1); |
|
457 |
qed "domain_0"; |
|
12199 | 458 |
Addsimps [domain_0]; |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
459 |
|
9180 | 460 |
Goal "domain(cons(<a,b>,r)) = cons(a, domain(r))"; |
461 |
by (Blast_tac 1); |
|
462 |
qed "domain_cons"; |
|
12199 | 463 |
Addsimps [domain_cons]; |
0 | 464 |
|
5321 | 465 |
Goal "domain(A Un B) = domain(A) Un domain(B)"; |
2877 | 466 |
by (Blast_tac 1); |
760 | 467 |
qed "domain_Un_eq"; |
12199 | 468 |
Addsimps [domain_Un_eq]; |
0 | 469 |
|
5321 | 470 |
Goal "domain(A Int B) <= domain(A) Int domain(B)"; |
2877 | 471 |
by (Blast_tac 1); |
760 | 472 |
qed "domain_Int_subset"; |
0 | 473 |
|
5321 | 474 |
Goal "domain(A) - domain(B) <= domain(A - B)"; |
2877 | 475 |
by (Blast_tac 1); |
1056 | 476 |
qed "domain_Diff_subset"; |
0 | 477 |
|
5321 | 478 |
Goal "domain(converse(r)) = range(r)"; |
2877 | 479 |
by (Blast_tac 1); |
760 | 480 |
qed "domain_converse"; |
12199 | 481 |
Addsimps [domain_converse]; |
685 | 482 |
|
12199 | 483 |
Goal "domain(UN x:A. B(x)) = (UN x:A. domain(B(x)))"; |
484 |
by (Blast_tac 1); |
|
485 |
qed "domain_UN"; |
|
486 |
||
487 |
Goal "domain(Union(A)) = (UN x:A. domain(x))"; |
|
488 |
by (Blast_tac 1); |
|
489 |
qed "domain_Union"; |
|
685 | 490 |
|
491 |
||
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
492 |
(** Range **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
493 |
|
5321 | 494 |
Goal "a:A ==> range(A*B) = B"; |
495 |
by (Blast_tac 1); |
|
496 |
qed "range_of_prod"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
497 |
|
9180 | 498 |
Goal "range(0) = 0"; |
499 |
by (Blast_tac 1); |
|
500 |
qed "range_0"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
501 |
|
9180 | 502 |
Goal "range(cons(<a,b>,r)) = cons(b, range(r))"; |
503 |
by (Blast_tac 1); |
|
504 |
qed "range_cons"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
505 |
|
5321 | 506 |
Goal "range(A Un B) = range(A) Un range(B)"; |
2877 | 507 |
by (Blast_tac 1); |
760 | 508 |
qed "range_Un_eq"; |
0 | 509 |
|
5321 | 510 |
Goal "range(A Int B) <= range(A) Int range(B)"; |
2877 | 511 |
by (Blast_tac 1); |
760 | 512 |
qed "range_Int_subset"; |
0 | 513 |
|
5321 | 514 |
Goal "range(A) - range(B) <= range(A - B)"; |
2877 | 515 |
by (Blast_tac 1); |
1056 | 516 |
qed "range_Diff_subset"; |
0 | 517 |
|
5321 | 518 |
Goal "range(converse(r)) = domain(r)"; |
2877 | 519 |
by (Blast_tac 1); |
760 | 520 |
qed "range_converse"; |
685 | 521 |
|
2469 | 522 |
Addsimps [range_0, range_cons, range_Un_eq, range_converse]; |
523 |
||
524 |
||
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
525 |
(** Field **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
526 |
|
9180 | 527 |
Goal "field(A*A) = A"; |
528 |
by (Blast_tac 1); |
|
529 |
qed "field_of_prod"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
530 |
|
9180 | 531 |
Goal "field(0) = 0"; |
532 |
by (Blast_tac 1); |
|
533 |
qed "field_0"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
534 |
|
9180 | 535 |
Goal "field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"; |
536 |
by (rtac equalityI 1); |
|
537 |
by (ALLGOALS Blast_tac) ; |
|
538 |
qed "field_cons"; |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
539 |
|
5321 | 540 |
Goal "field(A Un B) = field(A) Un field(B)"; |
2877 | 541 |
by (Blast_tac 1); |
760 | 542 |
qed "field_Un_eq"; |
0 | 543 |
|
5321 | 544 |
Goal "field(A Int B) <= field(A) Int field(B)"; |
2877 | 545 |
by (Blast_tac 1); |
760 | 546 |
qed "field_Int_subset"; |
0 | 547 |
|
5321 | 548 |
Goal "field(A) - field(B) <= field(A - B)"; |
2877 | 549 |
by (Blast_tac 1); |
1056 | 550 |
qed "field_Diff_subset"; |
0 | 551 |
|
5321 | 552 |
Goal "field(converse(r)) = field(r)"; |
2877 | 553 |
by (Blast_tac 1); |
2469 | 554 |
qed "field_converse"; |
555 |
||
556 |
Addsimps [field_0, field_cons, field_Un_eq, field_converse]; |
|
557 |
||
0 | 558 |
|
435 | 559 |
(** Image **) |
560 |
||
5321 | 561 |
Goal "r``0 = 0"; |
2877 | 562 |
by (Blast_tac 1); |
760 | 563 |
qed "image_0"; |
435 | 564 |
|
5321 | 565 |
Goal "r``(A Un B) = (r``A) Un (r``B)"; |
2877 | 566 |
by (Blast_tac 1); |
760 | 567 |
qed "image_Un"; |
435 | 568 |
|
5321 | 569 |
Goal "r``(A Int B) <= (r``A) Int (r``B)"; |
2877 | 570 |
by (Blast_tac 1); |
760 | 571 |
qed "image_Int_subset"; |
435 | 572 |
|
5321 | 573 |
Goal "(r Int A*A)``B <= (r``B) Int A"; |
2877 | 574 |
by (Blast_tac 1); |
760 | 575 |
qed "image_Int_square_subset"; |
435 | 576 |
|
5321 | 577 |
Goal "B<=A ==> (r Int A*A)``B = (r``B) Int A"; |
2877 | 578 |
by (Blast_tac 1); |
760 | 579 |
qed "image_Int_square"; |
435 | 580 |
|
2469 | 581 |
Addsimps [image_0, image_Un]; |
582 |
||
4840 | 583 |
(*Image laws for special relations*) |
5321 | 584 |
Goal "0``A = 0"; |
4840 | 585 |
by (Blast_tac 1); |
586 |
qed "image_0_left"; |
|
587 |
Addsimps [image_0_left]; |
|
588 |
||
5321 | 589 |
Goal "(r Un s)``A = (r``A) Un (s``A)"; |
4840 | 590 |
by (Blast_tac 1); |
591 |
qed "image_Un_left"; |
|
592 |
||
5321 | 593 |
Goal "(r Int s)``A <= (r``A) Int (s``A)"; |
4840 | 594 |
by (Blast_tac 1); |
595 |
qed "image_Int_subset_left"; |
|
596 |
||
435 | 597 |
|
598 |
(** Inverse Image **) |
|
599 |
||
5321 | 600 |
Goal "r-``0 = 0"; |
2877 | 601 |
by (Blast_tac 1); |
760 | 602 |
qed "vimage_0"; |
12199 | 603 |
Addsimps [vimage_0]; |
435 | 604 |
|
5321 | 605 |
Goal "r-``(A Un B) = (r-``A) Un (r-``B)"; |
2877 | 606 |
by (Blast_tac 1); |
760 | 607 |
qed "vimage_Un"; |
12199 | 608 |
Addsimps [vimage_Un]; |
435 | 609 |
|
5321 | 610 |
Goal "r-``(A Int B) <= (r-``A) Int (r-``B)"; |
2877 | 611 |
by (Blast_tac 1); |
760 | 612 |
qed "vimage_Int_subset"; |
435 | 613 |
|
12199 | 614 |
(*NOT suitable for rewriting*) |
615 |
Goal "f -``B = (UN y:B. f-``{y})"; |
|
616 |
by (Blast_tac 1); |
|
617 |
qed "vimage_eq_UN"; |
|
618 |
||
5321 | 619 |
Goalw [function_def] "function(f) ==> f-``(A Int B) = (f-``A) Int (f-``B)"; |
4660 | 620 |
by (Blast_tac 1); |
621 |
qed "function_vimage_Int"; |
|
622 |
||
5321 | 623 |
Goalw [function_def] "function(f) ==> f-``(A-B) = (f-``A) - (f-``B)"; |
4660 | 624 |
by (Blast_tac 1); |
625 |
qed "function_vimage_Diff"; |
|
626 |
||
5321 | 627 |
Goalw [function_def] "function(f) ==> f `` (f-`` A) <= A"; |
628 |
by (Blast_tac 1); |
|
629 |
qed "function_image_vimage"; |
|
630 |
||
631 |
Goal "(r Int A*A)-``B <= (r-``B) Int A"; |
|
2877 | 632 |
by (Blast_tac 1); |
760 | 633 |
qed "vimage_Int_square_subset"; |
435 | 634 |
|
5321 | 635 |
Goal "B<=A ==> (r Int A*A)-``B = (r-``B) Int A"; |
2877 | 636 |
by (Blast_tac 1); |
760 | 637 |
qed "vimage_Int_square"; |
435 | 638 |
|
2469 | 639 |
|
435 | 640 |
|
4840 | 641 |
(*Invese image laws for special relations*) |
5321 | 642 |
Goal "0-``A = 0"; |
4840 | 643 |
by (Blast_tac 1); |
644 |
qed "vimage_0_left"; |
|
645 |
Addsimps [vimage_0_left]; |
|
646 |
||
5321 | 647 |
Goal "(r Un s)-``A = (r-``A) Un (s-``A)"; |
4840 | 648 |
by (Blast_tac 1); |
649 |
qed "vimage_Un_left"; |
|
650 |
||
5321 | 651 |
Goal "(r Int s)-``A <= (r-``A) Int (s-``A)"; |
4840 | 652 |
by (Blast_tac 1); |
653 |
qed "vimage_Int_subset_left"; |
|
654 |
||
655 |
||
0 | 656 |
(** Converse **) |
657 |
||
5321 | 658 |
Goal "converse(A Un B) = converse(A) Un converse(B)"; |
2877 | 659 |
by (Blast_tac 1); |
760 | 660 |
qed "converse_Un"; |
0 | 661 |
|
5321 | 662 |
Goal "converse(A Int B) = converse(A) Int converse(B)"; |
2877 | 663 |
by (Blast_tac 1); |
760 | 664 |
qed "converse_Int"; |
0 | 665 |
|
5321 | 666 |
Goal "converse(A - B) = converse(A) - converse(B)"; |
2877 | 667 |
by (Blast_tac 1); |
1056 | 668 |
qed "converse_Diff"; |
0 | 669 |
|
5321 | 670 |
Goal "converse(UN x:A. B(x)) = (UN x:A. converse(B(x)))"; |
2877 | 671 |
by (Blast_tac 1); |
787 | 672 |
qed "converse_UN"; |
673 |
||
1652 | 674 |
(*Unfolding Inter avoids using excluded middle on A=0*) |
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5321
diff
changeset
|
675 |
Goalw [Inter_def] "converse(INT x:A. B(x)) = (INT x:A. converse(B(x)))"; |
2877 | 676 |
by (Blast_tac 1); |
1652 | 677 |
qed "converse_INT"; |
678 |
||
2469 | 679 |
Addsimps [converse_Un, converse_Int, converse_Diff, converse_UN, converse_INT]; |
680 |
||
198 | 681 |
(** Pow **) |
682 |
||
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
683 |
Goal "Pow(0) = {0}"; |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
684 |
by (Blast_tac 1); |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
685 |
qed "Pow_0"; |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
686 |
|
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
687 |
Goal "Pow (cons(a,A)) = Pow(A) Un {cons(a,X) . X: Pow(A)}"; |
6298 | 688 |
by (rtac equalityI 1); |
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
689 |
by Safe_tac; |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
690 |
by (etac swap 1); |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
691 |
by (res_inst_tac [("a", "x-{a}")] RepFun_eqI 1); |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
692 |
by (ALLGOALS Blast_tac); |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
693 |
qed "Pow_insert"; |
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
694 |
|
5321 | 695 |
Goal "Pow(A) Un Pow(B) <= Pow(A Un B)"; |
2877 | 696 |
by (Blast_tac 1); |
760 | 697 |
qed "Un_Pow_subset"; |
198 | 698 |
|
5321 | 699 |
Goal "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))"; |
2877 | 700 |
by (Blast_tac 1); |
760 | 701 |
qed "UN_Pow_subset"; |
198 | 702 |
|
5321 | 703 |
Goal "A <= Pow(Union(A))"; |
2877 | 704 |
by (Blast_tac 1); |
760 | 705 |
qed "subset_Pow_Union"; |
198 | 706 |
|
5321 | 707 |
Goal "Union(Pow(A)) = A"; |
2877 | 708 |
by (Blast_tac 1); |
760 | 709 |
qed "Union_Pow_eq"; |
198 | 710 |
|
5321 | 711 |
Goal "Pow(A Int B) = Pow(A) Int Pow(B)"; |
2877 | 712 |
by (Blast_tac 1); |
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
713 |
qed "Pow_Int_eq"; |
198 | 714 |
|
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
715 |
Goal "x:A ==> Pow(INT x:A. B(x)) = (INT x:A. Pow(B(x)))"; |
2877 | 716 |
by (Blast_tac 1); |
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
717 |
qed "Pow_INT_eq"; |
435 | 718 |
|
6288
694c9c1910e8
new theorems Pow_0 and Pow_insert; renamed other Pow theorems
paulson
parents:
6068
diff
changeset
|
719 |
Addsimps [Pow_0, Union_Pow_eq, Pow_Int_eq]; |
2469 | 720 |
|
839 | 721 |
(** RepFun **) |
722 |
||
5321 | 723 |
Goal "{f(x).x:A}=0 <-> A=0"; |
9303 | 724 |
by (Blast_tac 1); |
839 | 725 |
qed "RepFun_eq_0_iff"; |
12199 | 726 |
Addsimps [RepFun_eq_0_iff]; |
727 |
||
728 |
Goal "{c. x:A} = (if A=0 then 0 else {c})"; |
|
729 |
by Auto_tac; |
|
730 |
by (Blast_tac 1); |
|
731 |
qed "RepFun_constant"; |
|
732 |
Addsimps [RepFun_constant]; |
|
839 | 733 |
|
1611 | 734 |
(** Collect **) |
735 |
||
5321 | 736 |
Goal "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)"; |
2877 | 737 |
by (Blast_tac 1); |
1611 | 738 |
qed "Collect_Un"; |
739 |
||
5321 | 740 |
Goal "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)"; |
2877 | 741 |
by (Blast_tac 1); |
1611 | 742 |
qed "Collect_Int"; |
743 |
||
5321 | 744 |
Goal "Collect(A - B, P) = Collect(A,P) - Collect(B,P)"; |
2877 | 745 |
by (Blast_tac 1); |
1611 | 746 |
qed "Collect_Diff"; |
747 |
||
6068 | 748 |
Goal "{x:cons(a,B). P(x)} = \ |
749 |
\ (if P(a) then cons(a, {x:B. P(x)}) else {x:B. P(x)})"; |
|
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5067
diff
changeset
|
750 |
by (simp_tac (simpset() addsplits [split_if]) 1); |
2877 | 751 |
by (Blast_tac 1); |
1611 | 752 |
qed "Collect_cons"; |