author | wenzelm |
Mon, 06 Oct 1997 19:15:22 +0200 | |
changeset 3795 | e687069e7257 |
parent 2925 | b0ae2e13db93 |
child 4091 | 771b1f6422a8 |
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!*) |
2877 | 13 |
goal ZF.thy "{a} Un B = cons(a,B)"; |
14 |
by (Blast_tac 1); |
|
760 | 15 |
qed "cons_eq"; |
520 | 16 |
|
2877 | 17 |
goal ZF.thy "cons(a, cons(b, C)) = cons(b, cons(a, C))"; |
18 |
by (Blast_tac 1); |
|
760 | 19 |
qed "cons_commute"; |
0 | 20 |
|
2877 | 21 |
goal ZF.thy "!!B. a: B ==> cons(a,B) = B"; |
22 |
by (Blast_tac 1); |
|
760 | 23 |
qed "cons_absorb"; |
0 | 24 |
|
2877 | 25 |
goal ZF.thy "!!B. a: B ==> cons(a, B-{a}) = B"; |
26 |
by (Blast_tac 1); |
|
760 | 27 |
qed "cons_Diff"; |
0 | 28 |
|
2877 | 29 |
goal ZF.thy "!!C. [| a: C; ALL y:C. y=b |] ==> C = {b}"; |
30 |
by (Blast_tac 1); |
|
760 | 31 |
qed "equal_singleton_lemma"; |
0 | 32 |
val equal_singleton = ballI RSN (2,equal_singleton_lemma); |
33 |
||
34 |
||
35 |
(** Binary Intersection **) |
|
36 |
||
37 |
(*NOT an equality, but it seems to belong here...*) |
|
2877 | 38 |
goal ZF.thy "cons(a,B) Int C <= cons(a, B Int C)"; |
39 |
by (Blast_tac 1); |
|
760 | 40 |
qed "Int_cons"; |
0 | 41 |
|
2877 | 42 |
goal ZF.thy "A Int A = A"; |
43 |
by (Blast_tac 1); |
|
760 | 44 |
qed "Int_absorb"; |
0 | 45 |
|
2877 | 46 |
goal ZF.thy "A Int B = B Int A"; |
47 |
by (Blast_tac 1); |
|
760 | 48 |
qed "Int_commute"; |
0 | 49 |
|
2877 | 50 |
goal ZF.thy "(A Int B) Int C = A Int (B Int C)"; |
51 |
by (Blast_tac 1); |
|
760 | 52 |
qed "Int_assoc"; |
0 | 53 |
|
2877 | 54 |
goal ZF.thy "(A Un B) Int C = (A Int C) Un (B Int C)"; |
55 |
by (Blast_tac 1); |
|
760 | 56 |
qed "Int_Un_distrib"; |
0 | 57 |
|
2877 | 58 |
goal ZF.thy "A<=B <-> A Int B = A"; |
59 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 60 |
qed "subset_Int_iff"; |
0 | 61 |
|
2877 | 62 |
goal ZF.thy "A<=B <-> B Int A = A"; |
63 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 64 |
qed "subset_Int_iff2"; |
435 | 65 |
|
2877 | 66 |
goal ZF.thy "!!A B C. C<=A ==> (A-B) Int C = C-B"; |
67 |
by (Blast_tac 1); |
|
1035 | 68 |
qed "Int_Diff_eq"; |
69 |
||
0 | 70 |
(** Binary Union **) |
71 |
||
2877 | 72 |
goal ZF.thy "cons(a,B) Un C = cons(a, B Un C)"; |
73 |
by (Blast_tac 1); |
|
760 | 74 |
qed "Un_cons"; |
0 | 75 |
|
2877 | 76 |
goal ZF.thy "A Un A = A"; |
77 |
by (Blast_tac 1); |
|
760 | 78 |
qed "Un_absorb"; |
0 | 79 |
|
2877 | 80 |
goal ZF.thy "A Un B = B Un A"; |
81 |
by (Blast_tac 1); |
|
760 | 82 |
qed "Un_commute"; |
0 | 83 |
|
2877 | 84 |
goal ZF.thy "(A Un B) Un C = A Un (B Un C)"; |
85 |
by (Blast_tac 1); |
|
760 | 86 |
qed "Un_assoc"; |
0 | 87 |
|
2877 | 88 |
goal ZF.thy "(A Int B) Un C = (A Un C) Int (B Un C)"; |
89 |
by (Blast_tac 1); |
|
760 | 90 |
qed "Un_Int_distrib"; |
0 | 91 |
|
2877 | 92 |
goal ZF.thy "A<=B <-> A Un B = B"; |
93 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 94 |
qed "subset_Un_iff"; |
0 | 95 |
|
2877 | 96 |
goal ZF.thy "A<=B <-> B Un A = B"; |
97 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 98 |
qed "subset_Un_iff2"; |
435 | 99 |
|
0 | 100 |
(** Simple properties of Diff -- set difference **) |
101 |
||
2877 | 102 |
goal ZF.thy "A-A = 0"; |
103 |
by (Blast_tac 1); |
|
760 | 104 |
qed "Diff_cancel"; |
0 | 105 |
|
2877 | 106 |
goal ZF.thy "0-A = 0"; |
107 |
by (Blast_tac 1); |
|
760 | 108 |
qed "empty_Diff"; |
0 | 109 |
|
2877 | 110 |
goal ZF.thy "A-0 = A"; |
111 |
by (Blast_tac 1); |
|
760 | 112 |
qed "Diff_0"; |
0 | 113 |
|
2877 | 114 |
goal ZF.thy "A-B=0 <-> A<=B"; |
115 |
by (blast_tac (!claset addEs [equalityE]) 1); |
|
787 | 116 |
qed "Diff_eq_0_iff"; |
117 |
||
0 | 118 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
2877 | 119 |
goal ZF.thy "A - cons(a,B) = A - B - {a}"; |
120 |
by (Blast_tac 1); |
|
760 | 121 |
qed "Diff_cons"; |
0 | 122 |
|
123 |
(*NOT SUITABLE FOR REWRITING since {a} == cons(a,0)*) |
|
2877 | 124 |
goal ZF.thy "A - cons(a,B) = A - {a} - B"; |
125 |
by (Blast_tac 1); |
|
760 | 126 |
qed "Diff_cons2"; |
0 | 127 |
|
2877 | 128 |
goal ZF.thy "A Int (B-A) = 0"; |
129 |
by (Blast_tac 1); |
|
760 | 130 |
qed "Diff_disjoint"; |
0 | 131 |
|
2877 | 132 |
goal ZF.thy "!!A B. A<=B ==> A Un (B-A) = B"; |
133 |
by (Blast_tac 1); |
|
760 | 134 |
qed "Diff_partition"; |
0 | 135 |
|
2877 | 136 |
goal ZF.thy "A <= B Un (A - B)"; |
137 |
by (Blast_tac 1); |
|
1611 | 138 |
qed "subset_Un_Diff"; |
139 |
||
2877 | 140 |
goal ZF.thy "!!A B. [| A<=B; B<=C |] ==> B-(C-A) = A"; |
141 |
by (Blast_tac 1); |
|
760 | 142 |
qed "double_complement"; |
0 | 143 |
|
2877 | 144 |
goal ZF.thy "(A Un B) - (B-A) = A"; |
145 |
by (Blast_tac 1); |
|
760 | 146 |
qed "double_complement_Un"; |
268 | 147 |
|
2877 | 148 |
goal ZF.thy |
0 | 149 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; |
2877 | 150 |
by (Blast_tac 1); |
760 | 151 |
qed "Un_Int_crazy"; |
0 | 152 |
|
2877 | 153 |
goal ZF.thy "A - (B Un C) = (A-B) Int (A-C)"; |
154 |
by (Blast_tac 1); |
|
760 | 155 |
qed "Diff_Un"; |
0 | 156 |
|
2877 | 157 |
goal ZF.thy "A - (B Int C) = (A-B) Un (A-C)"; |
158 |
by (Blast_tac 1); |
|
760 | 159 |
qed "Diff_Int"; |
0 | 160 |
|
161 |
(*Halmos, Naive Set Theory, page 16.*) |
|
2877 | 162 |
goal ZF.thy "(A Int B) Un C = A Int (B Un C) <-> C<=A"; |
163 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 164 |
qed "Un_Int_assoc_iff"; |
0 | 165 |
|
166 |
||
167 |
(** Big Union and Intersection **) |
|
168 |
||
2877 | 169 |
goal ZF.thy "Union(cons(a,B)) = a Un Union(B)"; |
170 |
by (Blast_tac 1); |
|
760 | 171 |
qed "Union_cons"; |
0 | 172 |
|
2877 | 173 |
goal ZF.thy "Union(A Un B) = Union(A) Un Union(B)"; |
174 |
by (Blast_tac 1); |
|
760 | 175 |
qed "Union_Un_distrib"; |
0 | 176 |
|
2877 | 177 |
goal ZF.thy "Union(A Int B) <= Union(A) Int Union(B)"; |
178 |
by (Blast_tac 1); |
|
760 | 179 |
qed "Union_Int_subset"; |
435 | 180 |
|
2877 | 181 |
goal ZF.thy "Union(C) Int A = 0 <-> (ALL B:C. B Int A = 0)"; |
182 |
by (blast_tac (!claset addSEs [equalityE]) 1); |
|
760 | 183 |
qed "Union_disjoint"; |
0 | 184 |
|
2877 | 185 |
goalw ZF.thy [Inter_def] "Inter(0) = 0"; |
186 |
by (Blast_tac 1); |
|
1652 | 187 |
qed "Inter_0"; |
188 |
||
2877 | 189 |
goal ZF.thy "!!A B. [| z:A; z:B |] ==> Inter(A) Un Inter(B) <= Inter(A Int B)"; |
190 |
by (Blast_tac 1); |
|
1568 | 191 |
qed "Inter_Un_subset"; |
192 |
||
0 | 193 |
(* A good challenge: Inter is ill-behaved on the empty set *) |
2877 | 194 |
goal ZF.thy "!!A B. [| a:A; b:B |] ==> Inter(A Un B) = Inter(A) Int Inter(B)"; |
195 |
by (Blast_tac 1); |
|
760 | 196 |
qed "Inter_Un_distrib"; |
0 | 197 |
|
2877 | 198 |
goal ZF.thy "Union({b}) = b"; |
199 |
by (Blast_tac 1); |
|
760 | 200 |
qed "Union_singleton"; |
0 | 201 |
|
2877 | 202 |
goal ZF.thy "Inter({b}) = b"; |
203 |
by (Blast_tac 1); |
|
760 | 204 |
qed "Inter_singleton"; |
0 | 205 |
|
206 |
(** Unions and Intersections of Families **) |
|
207 |
||
2877 | 208 |
goal ZF.thy "Union(A) = (UN x:A. x)"; |
209 |
by (Blast_tac 1); |
|
760 | 210 |
qed "Union_eq_UN"; |
0 | 211 |
|
2877 | 212 |
goalw ZF.thy [Inter_def] "Inter(A) = (INT x:A. x)"; |
213 |
by (Blast_tac 1); |
|
760 | 214 |
qed "Inter_eq_INT"; |
0 | 215 |
|
2877 | 216 |
goal ZF.thy "(UN i:0. A(i)) = 0"; |
217 |
by (Blast_tac 1); |
|
760 | 218 |
qed "UN_0"; |
517 | 219 |
|
0 | 220 |
(*Halmos, Naive Set Theory, page 35.*) |
2877 | 221 |
goal ZF.thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; |
222 |
by (Blast_tac 1); |
|
760 | 223 |
qed "Int_UN_distrib"; |
0 | 224 |
|
2877 | 225 |
goal ZF.thy "!!A B. i:I ==> B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; |
226 |
by (Blast_tac 1); |
|
760 | 227 |
qed "Un_INT_distrib"; |
0 | 228 |
|
2877 | 229 |
goal ZF.thy |
0 | 230 |
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; |
2877 | 231 |
by (Blast_tac 1); |
760 | 232 |
qed "Int_UN_distrib2"; |
0 | 233 |
|
2877 | 234 |
goal ZF.thy |
0 | 235 |
"!!I J. [| i:I; j:J |] ==> \ |
236 |
\ (INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; |
|
2877 | 237 |
by (Blast_tac 1); |
760 | 238 |
qed "Un_INT_distrib2"; |
0 | 239 |
|
2877 | 240 |
goal ZF.thy "!!A. a: A ==> (UN y:A. c) = c"; |
241 |
by (Blast_tac 1); |
|
760 | 242 |
qed "UN_constant"; |
0 | 243 |
|
2877 | 244 |
goal ZF.thy "!!A. a: A ==> (INT y:A. c) = c"; |
245 |
by (Blast_tac 1); |
|
760 | 246 |
qed "INT_constant"; |
0 | 247 |
|
248 |
(** Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: |
|
249 |
Union of a family of unions **) |
|
250 |
||
2877 | 251 |
goal ZF.thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; |
252 |
by (Blast_tac 1); |
|
760 | 253 |
qed "UN_Un_distrib"; |
0 | 254 |
|
2877 | 255 |
goal ZF.thy |
0 | 256 |
"!!A B. i:I ==> \ |
192 | 257 |
\ (INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; |
2877 | 258 |
by (Blast_tac 1); |
760 | 259 |
qed "INT_Int_distrib"; |
0 | 260 |
|
2877 | 261 |
goal ZF.thy "(UN z:I Int J. A(z)) <= (UN z:I. A(z)) Int (UN z:J. A(z))"; |
262 |
by (Blast_tac 1); |
|
1784 | 263 |
qed "UN_Int_subset"; |
264 |
||
0 | 265 |
(** Devlin, page 12, exercise 5: Complements **) |
266 |
||
2877 | 267 |
goal ZF.thy "!!A B. i:I ==> B - (UN i:I. A(i)) = (INT i:I. B - A(i))"; |
268 |
by (Blast_tac 1); |
|
760 | 269 |
qed "Diff_UN"; |
0 | 270 |
|
2877 | 271 |
goal ZF.thy "!!A B. i:I ==> B - (INT i:I. A(i)) = (UN i:I. B - A(i))"; |
272 |
by (Blast_tac 1); |
|
760 | 273 |
qed "Diff_INT"; |
0 | 274 |
|
275 |
(** Unions and Intersections with General Sum **) |
|
276 |
||
1611 | 277 |
(*Not suitable for rewriting: LOOPS!*) |
2877 | 278 |
goal ZF.thy "Sigma(cons(a,B), C) = ({a}*C(a)) Un Sigma(B,C)"; |
279 |
by (Blast_tac 1); |
|
1611 | 280 |
qed "Sigma_cons1"; |
281 |
||
282 |
(*Not suitable for rewriting: LOOPS!*) |
|
2877 | 283 |
goal ZF.thy "A * cons(b,B) = A*{b} Un A*B"; |
284 |
by (Blast_tac 1); |
|
1611 | 285 |
qed "Sigma_cons2"; |
286 |
||
2877 | 287 |
goal ZF.thy "Sigma(succ(A), B) = ({A}*B(A)) Un Sigma(A,B)"; |
288 |
by (Blast_tac 1); |
|
1611 | 289 |
qed "Sigma_succ1"; |
290 |
||
2877 | 291 |
goal ZF.thy "A * succ(B) = A*{B} Un A*B"; |
292 |
by (Blast_tac 1); |
|
1611 | 293 |
qed "Sigma_succ2"; |
520 | 294 |
|
2877 | 295 |
goal ZF.thy "(SUM x:(UN y:A. C(y)). B(x)) = (UN y:A. SUM x:C(y). B(x))"; |
296 |
by (Blast_tac 1); |
|
760 | 297 |
qed "SUM_UN_distrib1"; |
182 | 298 |
|
2877 | 299 |
goal ZF.thy "(SUM i:I. UN j:J. C(i,j)) = (UN j:J. SUM i:I. C(i,j))"; |
300 |
by (Blast_tac 1); |
|
760 | 301 |
qed "SUM_UN_distrib2"; |
182 | 302 |
|
2877 | 303 |
goal ZF.thy "(SUM i:I Un J. C(i)) = (SUM i:I. C(i)) Un (SUM j:J. C(j))"; |
304 |
by (Blast_tac 1); |
|
760 | 305 |
qed "SUM_Un_distrib1"; |
0 | 306 |
|
2877 | 307 |
goal ZF.thy "(SUM i:I. A(i) Un B(i)) = (SUM i:I. A(i)) Un (SUM i:I. B(i))"; |
308 |
by (Blast_tac 1); |
|
760 | 309 |
qed "SUM_Un_distrib2"; |
0 | 310 |
|
685 | 311 |
(*First-order version of the above, for rewriting*) |
2877 | 312 |
goal ZF.thy "I * (A Un B) = I*A Un I*B"; |
1461 | 313 |
by (rtac SUM_Un_distrib2 1); |
760 | 314 |
qed "prod_Un_distrib2"; |
685 | 315 |
|
2877 | 316 |
goal ZF.thy "(SUM i:I Int J. C(i)) = (SUM i:I. C(i)) Int (SUM j:J. C(j))"; |
317 |
by (Blast_tac 1); |
|
760 | 318 |
qed "SUM_Int_distrib1"; |
0 | 319 |
|
2877 | 320 |
goal ZF.thy |
192 | 321 |
"(SUM i:I. A(i) Int B(i)) = (SUM i:I. A(i)) Int (SUM i:I. B(i))"; |
2877 | 322 |
by (Blast_tac 1); |
760 | 323 |
qed "SUM_Int_distrib2"; |
0 | 324 |
|
685 | 325 |
(*First-order version of the above, for rewriting*) |
2877 | 326 |
goal ZF.thy "I * (A Int B) = I*A Int I*B"; |
1461 | 327 |
by (rtac SUM_Int_distrib2 1); |
760 | 328 |
qed "prod_Int_distrib2"; |
685 | 329 |
|
192 | 330 |
(*Cf Aczel, Non-Well-Founded Sets, page 115*) |
2877 | 331 |
goal ZF.thy "(SUM i:I. A(i)) = (UN i:I. {i} * A(i))"; |
332 |
by (Blast_tac 1); |
|
760 | 333 |
qed "SUM_eq_UN"; |
192 | 334 |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
335 |
(** Domain **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
336 |
|
2877 | 337 |
qed_goal "domain_of_prod" ZF.thy "!!A B. b:B ==> domain(A*B) = A" |
338 |
(fn _ => [ Blast_tac 1 ]); |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
339 |
|
2877 | 340 |
qed_goal "domain_0" ZF.thy "domain(0) = 0" |
341 |
(fn _ => [ Blast_tac 1 ]); |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
342 |
|
2877 | 343 |
qed_goal "domain_cons" ZF.thy |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
344 |
"domain(cons(<a,b>,r)) = cons(a, domain(r))" |
2877 | 345 |
(fn _ => [ Blast_tac 1 ]); |
0 | 346 |
|
2877 | 347 |
goal ZF.thy "domain(A Un B) = domain(A) Un domain(B)"; |
348 |
by (Blast_tac 1); |
|
760 | 349 |
qed "domain_Un_eq"; |
0 | 350 |
|
2877 | 351 |
goal ZF.thy "domain(A Int B) <= domain(A) Int domain(B)"; |
352 |
by (Blast_tac 1); |
|
760 | 353 |
qed "domain_Int_subset"; |
0 | 354 |
|
2877 | 355 |
goal ZF.thy "domain(A) - domain(B) <= domain(A - B)"; |
356 |
by (Blast_tac 1); |
|
1056 | 357 |
qed "domain_Diff_subset"; |
0 | 358 |
|
2877 | 359 |
goal ZF.thy "domain(converse(r)) = range(r)"; |
360 |
by (Blast_tac 1); |
|
760 | 361 |
qed "domain_converse"; |
685 | 362 |
|
2469 | 363 |
Addsimps [domain_0, domain_cons, domain_Un_eq, domain_converse]; |
685 | 364 |
|
365 |
||
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
366 |
(** Range **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
367 |
|
2877 | 368 |
qed_goal "range_of_prod" ZF.thy |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
369 |
"!!a A B. a:A ==> range(A*B) = B" |
2877 | 370 |
(fn _ => [ Blast_tac 1 ]); |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
371 |
|
2877 | 372 |
qed_goal "range_0" ZF.thy "range(0) = 0" |
373 |
(fn _ => [ Blast_tac 1 ]); |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
374 |
|
2877 | 375 |
qed_goal "range_cons" ZF.thy |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
376 |
"range(cons(<a,b>,r)) = cons(b, range(r))" |
2877 | 377 |
(fn _ => [ Blast_tac 1 ]); |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
378 |
|
2877 | 379 |
goal ZF.thy "range(A Un B) = range(A) Un range(B)"; |
380 |
by (Blast_tac 1); |
|
760 | 381 |
qed "range_Un_eq"; |
0 | 382 |
|
2877 | 383 |
goal ZF.thy "range(A Int B) <= range(A) Int range(B)"; |
384 |
by (Blast_tac 1); |
|
760 | 385 |
qed "range_Int_subset"; |
0 | 386 |
|
2877 | 387 |
goal ZF.thy "range(A) - range(B) <= range(A - B)"; |
388 |
by (Blast_tac 1); |
|
1056 | 389 |
qed "range_Diff_subset"; |
0 | 390 |
|
2877 | 391 |
goal ZF.thy "range(converse(r)) = domain(r)"; |
392 |
by (Blast_tac 1); |
|
760 | 393 |
qed "range_converse"; |
685 | 394 |
|
2469 | 395 |
Addsimps [range_0, range_cons, range_Un_eq, range_converse]; |
396 |
||
397 |
||
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
398 |
(** Field **) |
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
399 |
|
2877 | 400 |
qed_goal "field_of_prod" ZF.thy "field(A*A) = A" |
401 |
(fn _ => [ Blast_tac 1 ]); |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
402 |
|
2877 | 403 |
qed_goal "field_0" ZF.thy "field(0) = 0" |
404 |
(fn _ => [ Blast_tac 1 ]); |
|
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
405 |
|
2877 | 406 |
qed_goal "field_cons" ZF.thy |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
407 |
"field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))" |
2877 | 408 |
(fn _ => [ rtac equalityI 1, ALLGOALS (Blast_tac) ]); |
536
5fbfa997f1b0
ZF/domrange/domain_of_prod, domain_empty, etc: moved to equalities.ML where
lcp
parents:
520
diff
changeset
|
409 |
|
2877 | 410 |
goal ZF.thy "field(A Un B) = field(A) Un field(B)"; |
411 |
by (Blast_tac 1); |
|
760 | 412 |
qed "field_Un_eq"; |
0 | 413 |
|
2877 | 414 |
goal ZF.thy "field(A Int B) <= field(A) Int field(B)"; |
415 |
by (Blast_tac 1); |
|
760 | 416 |
qed "field_Int_subset"; |
0 | 417 |
|
2877 | 418 |
goal ZF.thy "field(A) - field(B) <= field(A - B)"; |
419 |
by (Blast_tac 1); |
|
1056 | 420 |
qed "field_Diff_subset"; |
0 | 421 |
|
2877 | 422 |
goal ZF.thy "field(converse(r)) = field(r)"; |
423 |
by (Blast_tac 1); |
|
2469 | 424 |
qed "field_converse"; |
425 |
||
426 |
Addsimps [field_0, field_cons, field_Un_eq, field_converse]; |
|
427 |
||
0 | 428 |
|
435 | 429 |
(** Image **) |
430 |
||
2877 | 431 |
goal ZF.thy "r``0 = 0"; |
432 |
by (Blast_tac 1); |
|
760 | 433 |
qed "image_0"; |
435 | 434 |
|
2877 | 435 |
goal ZF.thy "r``(A Un B) = (r``A) Un (r``B)"; |
436 |
by (Blast_tac 1); |
|
760 | 437 |
qed "image_Un"; |
435 | 438 |
|
2877 | 439 |
goal ZF.thy "r``(A Int B) <= (r``A) Int (r``B)"; |
440 |
by (Blast_tac 1); |
|
760 | 441 |
qed "image_Int_subset"; |
435 | 442 |
|
2877 | 443 |
goal ZF.thy "(r Int A*A)``B <= (r``B) Int A"; |
444 |
by (Blast_tac 1); |
|
760 | 445 |
qed "image_Int_square_subset"; |
435 | 446 |
|
2877 | 447 |
goal ZF.thy "!!r. B<=A ==> (r Int A*A)``B = (r``B) Int A"; |
448 |
by (Blast_tac 1); |
|
760 | 449 |
qed "image_Int_square"; |
435 | 450 |
|
2469 | 451 |
Addsimps [image_0, image_Un]; |
452 |
||
435 | 453 |
|
454 |
(** Inverse Image **) |
|
455 |
||
2877 | 456 |
goal ZF.thy "r-``0 = 0"; |
457 |
by (Blast_tac 1); |
|
760 | 458 |
qed "vimage_0"; |
435 | 459 |
|
2877 | 460 |
goal ZF.thy "r-``(A Un B) = (r-``A) Un (r-``B)"; |
461 |
by (Blast_tac 1); |
|
760 | 462 |
qed "vimage_Un"; |
435 | 463 |
|
2877 | 464 |
goal ZF.thy "r-``(A Int B) <= (r-``A) Int (r-``B)"; |
465 |
by (Blast_tac 1); |
|
760 | 466 |
qed "vimage_Int_subset"; |
435 | 467 |
|
2877 | 468 |
goal ZF.thy "(r Int A*A)-``B <= (r-``B) Int A"; |
469 |
by (Blast_tac 1); |
|
760 | 470 |
qed "vimage_Int_square_subset"; |
435 | 471 |
|
2877 | 472 |
goal ZF.thy "!!r. B<=A ==> (r Int A*A)-``B = (r-``B) Int A"; |
473 |
by (Blast_tac 1); |
|
760 | 474 |
qed "vimage_Int_square"; |
435 | 475 |
|
2469 | 476 |
Addsimps [vimage_0, vimage_Un]; |
477 |
||
435 | 478 |
|
0 | 479 |
(** Converse **) |
480 |
||
2877 | 481 |
goal ZF.thy "converse(A Un B) = converse(A) Un converse(B)"; |
482 |
by (Blast_tac 1); |
|
760 | 483 |
qed "converse_Un"; |
0 | 484 |
|
2877 | 485 |
goal ZF.thy "converse(A Int B) = converse(A) Int converse(B)"; |
486 |
by (Blast_tac 1); |
|
760 | 487 |
qed "converse_Int"; |
0 | 488 |
|
2877 | 489 |
goal ZF.thy "converse(A - B) = converse(A) - converse(B)"; |
490 |
by (Blast_tac 1); |
|
1056 | 491 |
qed "converse_Diff"; |
0 | 492 |
|
2877 | 493 |
goal ZF.thy "converse(UN x:A. B(x)) = (UN x:A. converse(B(x)))"; |
494 |
by (Blast_tac 1); |
|
787 | 495 |
qed "converse_UN"; |
496 |
||
1652 | 497 |
(*Unfolding Inter avoids using excluded middle on A=0*) |
2877 | 498 |
goalw ZF.thy [Inter_def] "converse(INT x:A. B(x)) = (INT x:A. converse(B(x)))"; |
499 |
by (Blast_tac 1); |
|
1652 | 500 |
qed "converse_INT"; |
501 |
||
2469 | 502 |
Addsimps [converse_Un, converse_Int, converse_Diff, converse_UN, converse_INT]; |
503 |
||
198 | 504 |
(** Pow **) |
505 |
||
2877 | 506 |
goal ZF.thy "Pow(A) Un Pow(B) <= Pow(A Un B)"; |
507 |
by (Blast_tac 1); |
|
760 | 508 |
qed "Un_Pow_subset"; |
198 | 509 |
|
2877 | 510 |
goal ZF.thy "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))"; |
511 |
by (Blast_tac 1); |
|
760 | 512 |
qed "UN_Pow_subset"; |
198 | 513 |
|
2877 | 514 |
goal ZF.thy "A <= Pow(Union(A))"; |
515 |
by (Blast_tac 1); |
|
760 | 516 |
qed "subset_Pow_Union"; |
198 | 517 |
|
2877 | 518 |
goal ZF.thy "Union(Pow(A)) = A"; |
519 |
by (Blast_tac 1); |
|
760 | 520 |
qed "Union_Pow_eq"; |
198 | 521 |
|
2877 | 522 |
goal ZF.thy "Pow(A Int B) = Pow(A) Int Pow(B)"; |
523 |
by (Blast_tac 1); |
|
760 | 524 |
qed "Int_Pow_eq"; |
198 | 525 |
|
2877 | 526 |
goal ZF.thy "!!x A. x:A ==> (INT x:A. Pow(B(x))) = Pow(INT x:A. B(x))"; |
527 |
by (Blast_tac 1); |
|
760 | 528 |
qed "INT_Pow_subset"; |
435 | 529 |
|
2469 | 530 |
Addsimps [Union_Pow_eq, Int_Pow_eq]; |
531 |
||
839 | 532 |
(** RepFun **) |
533 |
||
2877 | 534 |
goal ZF.thy "{f(x).x:A}=0 <-> A=0"; |
2925 | 535 |
(*blast_tac takes too long to find a good depth*) |
536 |
by (Blast.depth_tac (!claset addSEs [equalityE]) 10 1); |
|
839 | 537 |
qed "RepFun_eq_0_iff"; |
538 |
||
1611 | 539 |
(** Collect **) |
540 |
||
2877 | 541 |
goal ZF.thy "Collect(A Un B, P) = Collect(A,P) Un Collect(B,P)"; |
542 |
by (Blast_tac 1); |
|
1611 | 543 |
qed "Collect_Un"; |
544 |
||
2877 | 545 |
goal ZF.thy "Collect(A Int B, P) = Collect(A,P) Int Collect(B,P)"; |
546 |
by (Blast_tac 1); |
|
1611 | 547 |
qed "Collect_Int"; |
548 |
||
2877 | 549 |
goal ZF.thy "Collect(A - B, P) = Collect(A,P) - Collect(B,P)"; |
550 |
by (Blast_tac 1); |
|
1611 | 551 |
qed "Collect_Diff"; |
552 |
||
2877 | 553 |
goal ZF.thy "{x:cons(a,B). P(x)} = if(P(a), cons(a, {x:B. P(x)}), {x:B. P(x)})"; |
2469 | 554 |
by (simp_tac (!simpset setloop split_tac [expand_if]) 1); |
2877 | 555 |
by (Blast_tac 1); |
1611 | 556 |
qed "Collect_cons"; |
557 |