| author | paulson |
| Fri, 03 Apr 1998 11:22:51 +0200 | |
| changeset 4775 | 66b1a7c42d94 |
| parent 4771 | f1bacbbe02a8 |
| child 4882 | 379314255a04 |
| 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 |
||
|
1754
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1748
diff
changeset
|
11 |
AddSIs [equalityI]; |
|
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
berghofe
parents:
1748
diff
changeset
|
12 |
|
| 1548 | 13 |
section "{}";
|
14 |
||
| 4059 | 15 |
goal thy "{x. False} = {}";
|
| 2891 | 16 |
by (Blast_tac 1); |
| 1531 | 17 |
qed "Collect_False_empty"; |
18 |
Addsimps [Collect_False_empty]; |
|
19 |
||
| 4059 | 20 |
goal thy "(A <= {}) = (A = {})";
|
| 2891 | 21 |
by (Blast_tac 1); |
| 1531 | 22 |
qed "subset_empty"; |
23 |
Addsimps [subset_empty]; |
|
24 |
||
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
25 |
goalw thy [psubset_def] "~ (A < {})";
|
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
26 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
27 |
qed "not_psubset_empty"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
28 |
AddIffs [not_psubset_empty]; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
29 |
|
| 4748 | 30 |
goal thy "{x. P x | Q x} = {x. P x} Un {x. Q x}";
|
31 |
by (Blast_tac 1); |
|
32 |
qed "Collect_disj_eq"; |
|
33 |
||
34 |
goal thy "{x. P x & Q x} = {x. P x} Int {x. Q x}";
|
|
35 |
by (Blast_tac 1); |
|
36 |
qed "Collect_conj_eq"; |
|
37 |
||
38 |
||
| 1548 | 39 |
section "insert"; |
| 923 | 40 |
|
| 1531 | 41 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a {}*)
|
| 4059 | 42 |
goal thy "insert a A = {a} Un A";
|
| 2891 | 43 |
by (Blast_tac 1); |
| 1531 | 44 |
qed "insert_is_Un"; |
45 |
||
| 4059 | 46 |
goal thy "insert a A ~= {}";
|
| 4089 | 47 |
by (blast_tac (claset() addEs [equalityCE]) 1); |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
48 |
qed"insert_not_empty"; |
| 1531 | 49 |
Addsimps[insert_not_empty]; |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
50 |
|
|
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
51 |
bind_thm("empty_not_insert",insert_not_empty RS not_sym);
|
| 1531 | 52 |
Addsimps[empty_not_insert]; |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
53 |
|
| 4059 | 54 |
goal thy "!!a. a:A ==> insert a A = A"; |
| 2891 | 55 |
by (Blast_tac 1); |
| 923 | 56 |
qed "insert_absorb"; |
| 4605 | 57 |
(* Addsimps [insert_absorb] causes recursive (ie quadtratic) calls |
58 |
in case of nested inserts! |
|
59 |
*) |
|
| 923 | 60 |
|
| 4059 | 61 |
goal thy "insert x (insert x A) = insert x A"; |
| 2891 | 62 |
by (Blast_tac 1); |
| 1531 | 63 |
qed "insert_absorb2"; |
64 |
Addsimps [insert_absorb2]; |
|
65 |
||
| 4059 | 66 |
goal thy "insert x (insert y A) = insert y (insert x A)"; |
| 2891 | 67 |
by (Blast_tac 1); |
| 1879 | 68 |
qed "insert_commute"; |
69 |
||
| 4059 | 70 |
goal thy "(insert x A <= B) = (x:B & A <= B)"; |
| 2891 | 71 |
by (Blast_tac 1); |
| 923 | 72 |
qed "insert_subset"; |
| 1531 | 73 |
Addsimps[insert_subset]; |
74 |
||
| 4059 | 75 |
goal thy "!!a. insert a A ~= insert a B ==> A ~= B"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
76 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
77 |
qed "insert_lim"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
78 |
|
| 1531 | 79 |
(* use new B rather than (A-{a}) to avoid infinite unfolding *)
|
| 4059 | 80 |
goal thy "!!a. a:A ==> ? B. A = insert a B & a ~: B"; |
| 1553 | 81 |
by (res_inst_tac [("x","A-{a}")] exI 1);
|
| 2891 | 82 |
by (Blast_tac 1); |
| 1531 | 83 |
qed "mk_disjoint_insert"; |
| 923 | 84 |
|
| 4059 | 85 |
goal thy |
|
1843
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
paulson
parents:
1786
diff
changeset
|
86 |
"!!A. A~={} ==> (UN x:A. insert a (B x)) = insert a (UN x:A. B x)";
|
| 2891 | 87 |
by (Blast_tac 1); |
|
1843
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
paulson
parents:
1786
diff
changeset
|
88 |
qed "UN_insert_distrib"; |
|
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
paulson
parents:
1786
diff
changeset
|
89 |
|
| 1660 | 90 |
section "``"; |
| 923 | 91 |
|
| 4059 | 92 |
goal thy "f``{} = {}";
|
| 2891 | 93 |
by (Blast_tac 1); |
| 923 | 94 |
qed "image_empty"; |
| 1531 | 95 |
Addsimps[image_empty]; |
| 923 | 96 |
|
| 4059 | 97 |
goal thy "f``insert a B = insert (f a) (f``B)"; |
| 2891 | 98 |
by (Blast_tac 1); |
| 923 | 99 |
qed "image_insert"; |
| 1531 | 100 |
Addsimps[image_insert]; |
| 923 | 101 |
|
| 4059 | 102 |
goal thy "(f `` (UNION A B)) = (UN x:A.(f `` (B x)))"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
103 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
104 |
qed "image_UNION"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
105 |
|
| 4059 | 106 |
goal thy "(%x. x) `` Y = Y"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
107 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
108 |
qed "image_id"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
109 |
|
| 4059 | 110 |
goal thy "f``(g``A) = (%x. f (g x)) `` A"; |
| 3457 | 111 |
by (Blast_tac 1); |
| 4059 | 112 |
qed "image_image"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
113 |
|
| 4059 | 114 |
goal thy "!!x. x:A ==> insert (f x) (f``A) = f``A"; |
| 2891 | 115 |
by (Blast_tac 1); |
| 1884 | 116 |
qed "insert_image"; |
117 |
Addsimps [insert_image]; |
|
118 |
||
| 4059 | 119 |
goal thy "(f``A = {}) = (A = {})";
|
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
120 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
|
3415
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents:
3384
diff
changeset
|
121 |
qed "image_is_empty"; |
|
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents:
3384
diff
changeset
|
122 |
AddIffs [image_is_empty]; |
|
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
nipkow
parents:
3384
diff
changeset
|
123 |
|
| 4059 | 124 |
goalw thy [image_def] |
| 1763 | 125 |
"(%x. if P x then f x else g x) `` S \ |
| 4200 | 126 |
\ = (f `` (S Int {x. P x})) Un (g `` (S Int {x. ~(P x)}))";
|
| 4686 | 127 |
by (Simp_tac 1); |
| 2891 | 128 |
by (Blast_tac 1); |
| 1748 | 129 |
qed "if_image_distrib"; |
130 |
Addsimps[if_image_distrib]; |
|
131 |
||
| 4136 | 132 |
val prems= goal thy "[|M = N; !!x. x:N ==> f x = g x|] ==> f``M = g``N"; |
133 |
by (rtac set_ext 1); |
|
134 |
by (simp_tac (simpset() addsimps image_def::prems) 1); |
|
135 |
qed "image_cong"; |
|
136 |
||
| 1748 | 137 |
|
| 1548 | 138 |
section "Int"; |
| 923 | 139 |
|
| 4059 | 140 |
goal thy "A Int A = A"; |
| 2891 | 141 |
by (Blast_tac 1); |
| 923 | 142 |
qed "Int_absorb"; |
| 1531 | 143 |
Addsimps[Int_absorb]; |
| 923 | 144 |
|
| 4609 | 145 |
goal thy " A Int (A Int B) = A Int B"; |
146 |
by (Blast_tac 1); |
|
147 |
qed "Int_left_absorb"; |
|
148 |
||
| 4059 | 149 |
goal thy "A Int B = B Int A"; |
| 2891 | 150 |
by (Blast_tac 1); |
| 923 | 151 |
qed "Int_commute"; |
152 |
||
| 4609 | 153 |
goal thy "A Int (B Int C) = B Int (A Int C)"; |
154 |
by (Blast_tac 1); |
|
155 |
qed "Int_left_commute"; |
|
156 |
||
| 4059 | 157 |
goal thy "(A Int B) Int C = A Int (B Int C)"; |
| 2891 | 158 |
by (Blast_tac 1); |
| 923 | 159 |
qed "Int_assoc"; |
160 |
||
| 4609 | 161 |
(*Intersection is an AC-operator*) |
162 |
val Int_ac = [Int_assoc, Int_left_absorb, Int_commute, Int_left_commute]; |
|
163 |
||
| 4662 | 164 |
goal thy "!!A B. B<=A ==> A Int B = B"; |
165 |
by (Blast_tac 1); |
|
166 |
qed "Int_absorb1"; |
|
167 |
||
168 |
goal thy "!!A B. A<=B ==> A Int B = A"; |
|
169 |
by (Blast_tac 1); |
|
170 |
qed "Int_absorb2"; |
|
171 |
||
| 4059 | 172 |
goal thy "{} Int B = {}";
|
| 2891 | 173 |
by (Blast_tac 1); |
| 923 | 174 |
qed "Int_empty_left"; |
| 1531 | 175 |
Addsimps[Int_empty_left]; |
| 923 | 176 |
|
| 4059 | 177 |
goal thy "A Int {} = {}";
|
| 2891 | 178 |
by (Blast_tac 1); |
| 923 | 179 |
qed "Int_empty_right"; |
| 1531 | 180 |
Addsimps[Int_empty_right]; |
181 |
||
| 4059 | 182 |
goal thy "(A Int B = {}) = (A <= Compl B)";
|
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
183 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
| 3356 | 184 |
qed "disjoint_eq_subset_Compl"; |
185 |
||
| 4059 | 186 |
goal thy "UNIV Int B = B"; |
| 2891 | 187 |
by (Blast_tac 1); |
| 1531 | 188 |
qed "Int_UNIV_left"; |
189 |
Addsimps[Int_UNIV_left]; |
|
190 |
||
| 4059 | 191 |
goal thy "A Int UNIV = A"; |
| 2891 | 192 |
by (Blast_tac 1); |
| 1531 | 193 |
qed "Int_UNIV_right"; |
194 |
Addsimps[Int_UNIV_right]; |
|
| 923 | 195 |
|
| 4634 | 196 |
goal thy "A Int B = Inter{A,B}";
|
197 |
by (Blast_tac 1); |
|
198 |
qed "Int_eq_Inter"; |
|
199 |
||
| 4059 | 200 |
goal thy "A Int (B Un C) = (A Int B) Un (A Int C)"; |
| 2891 | 201 |
by (Blast_tac 1); |
| 923 | 202 |
qed "Int_Un_distrib"; |
203 |
||
| 4059 | 204 |
goal thy "(B Un C) Int A = (B Int A) Un (C Int A)"; |
| 2891 | 205 |
by (Blast_tac 1); |
| 1618 | 206 |
qed "Int_Un_distrib2"; |
207 |
||
| 4059 | 208 |
goal thy "(A<=B) = (A Int B = A)"; |
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
209 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
| 923 | 210 |
qed "subset_Int_eq"; |
211 |
||
| 4059 | 212 |
goal thy "(A Int B = UNIV) = (A = UNIV & B = UNIV)"; |
| 4089 | 213 |
by (blast_tac (claset() addEs [equalityCE]) 1); |
| 1531 | 214 |
qed "Int_UNIV"; |
215 |
Addsimps[Int_UNIV]; |
|
216 |
||
| 1548 | 217 |
section "Un"; |
| 923 | 218 |
|
| 4059 | 219 |
goal thy "A Un A = A"; |
| 2891 | 220 |
by (Blast_tac 1); |
| 923 | 221 |
qed "Un_absorb"; |
| 1531 | 222 |
Addsimps[Un_absorb]; |
| 923 | 223 |
|
| 4059 | 224 |
goal thy " A Un (A Un B) = A Un B"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
225 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
226 |
qed "Un_left_absorb"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
227 |
|
| 4059 | 228 |
goal thy "A Un B = B Un A"; |
| 2891 | 229 |
by (Blast_tac 1); |
| 923 | 230 |
qed "Un_commute"; |
231 |
||
| 4609 | 232 |
goal thy "A Un (B Un C) = B Un (A Un C)"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
233 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
234 |
qed "Un_left_commute"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
235 |
|
| 4059 | 236 |
goal thy "(A Un B) Un C = A Un (B Un C)"; |
| 2891 | 237 |
by (Blast_tac 1); |
| 923 | 238 |
qed "Un_assoc"; |
239 |
||
| 4609 | 240 |
(*Union is an AC-operator*) |
241 |
val Un_ac = [Un_assoc, Un_left_absorb, Un_commute, Un_left_commute]; |
|
242 |
||
| 4662 | 243 |
goal thy "!!A B. A<=B ==> A Un B = B"; |
244 |
by (Blast_tac 1); |
|
245 |
qed "Un_absorb1"; |
|
246 |
||
247 |
goal thy "!!A B. B<=A ==> A Un B = A"; |
|
248 |
by (Blast_tac 1); |
|
249 |
qed "Un_absorb2"; |
|
250 |
||
| 4059 | 251 |
goal thy "{} Un B = B";
|
| 2891 | 252 |
by (Blast_tac 1); |
| 923 | 253 |
qed "Un_empty_left"; |
| 1531 | 254 |
Addsimps[Un_empty_left]; |
| 923 | 255 |
|
| 4059 | 256 |
goal thy "A Un {} = A";
|
| 2891 | 257 |
by (Blast_tac 1); |
| 923 | 258 |
qed "Un_empty_right"; |
| 1531 | 259 |
Addsimps[Un_empty_right]; |
260 |
||
| 4059 | 261 |
goal thy "UNIV Un B = UNIV"; |
| 2891 | 262 |
by (Blast_tac 1); |
| 1531 | 263 |
qed "Un_UNIV_left"; |
264 |
Addsimps[Un_UNIV_left]; |
|
265 |
||
| 4059 | 266 |
goal thy "A Un UNIV = UNIV"; |
| 2891 | 267 |
by (Blast_tac 1); |
| 1531 | 268 |
qed "Un_UNIV_right"; |
269 |
Addsimps[Un_UNIV_right]; |
|
| 923 | 270 |
|
| 4634 | 271 |
goal thy "A Un B = Union{A,B}";
|
272 |
by (Blast_tac 1); |
|
273 |
qed "Un_eq_Union"; |
|
274 |
||
| 4059 | 275 |
goal thy "(insert a B) Un C = insert a (B Un C)"; |
| 2891 | 276 |
by (Blast_tac 1); |
| 923 | 277 |
qed "Un_insert_left"; |
|
3384
5ef99c94e1fb
Now Un_insert_left, Un_insert_right are default rewrite rules
paulson
parents:
3356
diff
changeset
|
278 |
Addsimps[Un_insert_left]; |
| 923 | 279 |
|
| 4059 | 280 |
goal thy "A Un (insert a B) = insert a (A Un B)"; |
| 2891 | 281 |
by (Blast_tac 1); |
| 1917 | 282 |
qed "Un_insert_right"; |
|
3384
5ef99c94e1fb
Now Un_insert_left, Un_insert_right are default rewrite rules
paulson
parents:
3356
diff
changeset
|
283 |
Addsimps[Un_insert_right]; |
| 1917 | 284 |
|
| 4059 | 285 |
goal thy "(insert a B) Int C = (if a:C then insert a (B Int C) \ |
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
286 |
\ else B Int C)"; |
| 4686 | 287 |
by (Simp_tac 1); |
| 3356 | 288 |
by (Blast_tac 1); |
289 |
qed "Int_insert_left"; |
|
290 |
||
| 4059 | 291 |
goal thy "A Int (insert a B) = (if a:A then insert a (A Int B) \ |
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
292 |
\ else A Int B)"; |
| 4686 | 293 |
by (Simp_tac 1); |
| 3356 | 294 |
by (Blast_tac 1); |
295 |
qed "Int_insert_right"; |
|
296 |
||
| 4609 | 297 |
goal thy "A Un (B Int C) = (A Un B) Int (A Un C)"; |
| 2891 | 298 |
by (Blast_tac 1); |
| 923 | 299 |
qed "Un_Int_distrib"; |
300 |
||
| 4609 | 301 |
goal thy "(B Int C) Un A = (B Un A) Int (C Un A)"; |
302 |
by (Blast_tac 1); |
|
303 |
qed "Un_Int_distrib2"; |
|
304 |
||
| 4059 | 305 |
goal thy |
| 923 | 306 |
"(A Int B) Un (B Int C) Un (C Int A) = (A Un B) Int (B Un C) Int (C Un A)"; |
| 2891 | 307 |
by (Blast_tac 1); |
| 923 | 308 |
qed "Un_Int_crazy"; |
309 |
||
| 4059 | 310 |
goal thy "(A<=B) = (A Un B = B)"; |
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
311 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
| 923 | 312 |
qed "subset_Un_eq"; |
313 |
||
| 4059 | 314 |
goal thy "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)";
|
| 2891 | 315 |
by (Blast_tac 1); |
| 923 | 316 |
qed "subset_insert_iff"; |
317 |
||
| 4059 | 318 |
goal thy "(A Un B = {}) = (A = {} & B = {})";
|
| 4089 | 319 |
by (blast_tac (claset() addEs [equalityCE]) 1); |
| 923 | 320 |
qed "Un_empty"; |
| 1531 | 321 |
Addsimps[Un_empty]; |
| 923 | 322 |
|
| 1548 | 323 |
section "Compl"; |
| 923 | 324 |
|
| 4059 | 325 |
goal thy "A Int Compl(A) = {}";
|
| 2891 | 326 |
by (Blast_tac 1); |
| 923 | 327 |
qed "Compl_disjoint"; |
| 1531 | 328 |
Addsimps[Compl_disjoint]; |
| 923 | 329 |
|
| 4059 | 330 |
goal thy "A Un Compl(A) = UNIV"; |
| 2891 | 331 |
by (Blast_tac 1); |
| 923 | 332 |
qed "Compl_partition"; |
333 |
||
| 4059 | 334 |
goal thy "Compl(Compl(A)) = A"; |
| 2891 | 335 |
by (Blast_tac 1); |
| 923 | 336 |
qed "double_complement"; |
| 1531 | 337 |
Addsimps[double_complement]; |
| 923 | 338 |
|
| 4059 | 339 |
goal thy "Compl(A Un B) = Compl(A) Int Compl(B)"; |
| 2891 | 340 |
by (Blast_tac 1); |
| 923 | 341 |
qed "Compl_Un"; |
342 |
||
| 4059 | 343 |
goal thy "Compl(A Int B) = Compl(A) Un Compl(B)"; |
| 2891 | 344 |
by (Blast_tac 1); |
| 923 | 345 |
qed "Compl_Int"; |
346 |
||
| 4059 | 347 |
goal thy "Compl(UN x:A. B(x)) = (INT x:A. Compl(B(x)))"; |
| 2891 | 348 |
by (Blast_tac 1); |
| 923 | 349 |
qed "Compl_UN"; |
350 |
||
| 4059 | 351 |
goal thy "Compl(INT x:A. B(x)) = (UN x:A. Compl(B(x)))"; |
| 2891 | 352 |
by (Blast_tac 1); |
| 923 | 353 |
qed "Compl_INT"; |
354 |
||
| 4615 | 355 |
Addsimps [Compl_Un, Compl_Int, Compl_UN, Compl_INT]; |
356 |
||
| 923 | 357 |
(*Halmos, Naive Set Theory, page 16.*) |
358 |
||
| 4059 | 359 |
goal thy "((A Int B) Un C = A Int (B Un C)) = (C<=A)"; |
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
360 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
| 923 | 361 |
qed "Un_Int_assoc_eq"; |
362 |
||
363 |
||
| 1548 | 364 |
section "Union"; |
| 923 | 365 |
|
| 4059 | 366 |
goal thy "Union({}) = {}";
|
| 2891 | 367 |
by (Blast_tac 1); |
| 923 | 368 |
qed "Union_empty"; |
| 1531 | 369 |
Addsimps[Union_empty]; |
370 |
||
| 4059 | 371 |
goal thy "Union(UNIV) = UNIV"; |
| 2891 | 372 |
by (Blast_tac 1); |
| 1531 | 373 |
qed "Union_UNIV"; |
374 |
Addsimps[Union_UNIV]; |
|
| 923 | 375 |
|
| 4059 | 376 |
goal thy "Union(insert a B) = a Un Union(B)"; |
| 2891 | 377 |
by (Blast_tac 1); |
| 923 | 378 |
qed "Union_insert"; |
| 1531 | 379 |
Addsimps[Union_insert]; |
| 923 | 380 |
|
| 4059 | 381 |
goal thy "Union(A Un B) = Union(A) Un Union(B)"; |
| 2891 | 382 |
by (Blast_tac 1); |
| 923 | 383 |
qed "Union_Un_distrib"; |
| 1531 | 384 |
Addsimps[Union_Un_distrib]; |
| 923 | 385 |
|
| 4059 | 386 |
goal thy "Union(A Int B) <= Union(A) Int Union(B)"; |
| 2891 | 387 |
by (Blast_tac 1); |
| 923 | 388 |
qed "Union_Int_subset"; |
389 |
||
| 4059 | 390 |
goal thy "(Union M = {}) = (! A : M. A = {})";
|
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
391 |
by (blast_tac (claset() addEs [equalityCE]) 1); |
|
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
392 |
qed "Union_empty_conv"; |
| 4003 | 393 |
AddIffs [Union_empty_conv]; |
394 |
||
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
395 |
goal thy "(Union(C) Int A = {}) = (! B:C. B Int A = {})";
|
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
396 |
by (blast_tac (claset() addSEs [equalityCE]) 1); |
| 923 | 397 |
qed "Union_disjoint"; |
398 |
||
| 1548 | 399 |
section "Inter"; |
400 |
||
| 4059 | 401 |
goal thy "Inter({}) = UNIV";
|
| 2891 | 402 |
by (Blast_tac 1); |
| 1531 | 403 |
qed "Inter_empty"; |
404 |
Addsimps[Inter_empty]; |
|
405 |
||
| 4059 | 406 |
goal thy "Inter(UNIV) = {}";
|
| 2891 | 407 |
by (Blast_tac 1); |
| 1531 | 408 |
qed "Inter_UNIV"; |
409 |
Addsimps[Inter_UNIV]; |
|
410 |
||
| 4059 | 411 |
goal thy "Inter(insert a B) = a Int Inter(B)"; |
| 2891 | 412 |
by (Blast_tac 1); |
| 1531 | 413 |
qed "Inter_insert"; |
414 |
Addsimps[Inter_insert]; |
|
415 |
||
| 4059 | 416 |
goal thy "Inter(A) Un Inter(B) <= Inter(A Int B)"; |
| 2891 | 417 |
by (Blast_tac 1); |
|
1564
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
paulson
parents:
1553
diff
changeset
|
418 |
qed "Inter_Un_subset"; |
| 1531 | 419 |
|
| 4059 | 420 |
goal thy "Inter(A Un B) = Inter(A) Int Inter(B)"; |
| 2891 | 421 |
by (Blast_tac 1); |
| 923 | 422 |
qed "Inter_Un_distrib"; |
423 |
||
| 1548 | 424 |
section "UN and INT"; |
| 923 | 425 |
|
426 |
(*Basic identities*) |
|
427 |
||
| 4200 | 428 |
val not_empty = prove_goal Set.thy "(A ~= {}) = (? x. x:A)" (K [Blast_tac 1]);
|
| 4136 | 429 |
(*Addsimps[not_empty];*) |
430 |
||
| 4059 | 431 |
goal thy "(UN x:{}. B x) = {}";
|
| 2891 | 432 |
by (Blast_tac 1); |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
433 |
qed "UN_empty"; |
| 1531 | 434 |
Addsimps[UN_empty]; |
435 |
||
| 4059 | 436 |
goal thy "(UN x:A. {}) = {}";
|
| 3457 | 437 |
by (Blast_tac 1); |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
438 |
qed "UN_empty2"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
439 |
Addsimps[UN_empty2]; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
440 |
|
| 4645 | 441 |
goal thy "(UN x:A. {x}) = A";
|
442 |
by (Blast_tac 1); |
|
443 |
qed "UN_singleton"; |
|
444 |
Addsimps [UN_singleton]; |
|
445 |
||
| 4634 | 446 |
goal thy "!!k. k:I ==> A k Un (UN i:I. A i) = (UN i:I. A i)"; |
447 |
by (Blast_tac 1); |
|
448 |
qed "UN_absorb"; |
|
449 |
||
| 4059 | 450 |
goal thy "(INT x:{}. B x) = UNIV";
|
| 2891 | 451 |
by (Blast_tac 1); |
| 1531 | 452 |
qed "INT_empty"; |
453 |
Addsimps[INT_empty]; |
|
454 |
||
| 4634 | 455 |
goal thy "!!k. k:I ==> A k Int (INT i:I. A i) = (INT i:I. A i)"; |
456 |
by (Blast_tac 1); |
|
457 |
qed "INT_absorb"; |
|
458 |
||
| 4059 | 459 |
goal thy "(UN x:insert a A. B x) = B a Un UNION A B"; |
| 2891 | 460 |
by (Blast_tac 1); |
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
461 |
qed "UN_insert"; |
| 1531 | 462 |
Addsimps[UN_insert]; |
463 |
||
| 4059 | 464 |
goal thy "(UN i: A Un B. M i) = ((UN i: A. M i) Un (UN i:B. M i))"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
465 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
466 |
qed "UN_Un"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
467 |
|
| 4771 | 468 |
goal thy "(UN x : (UN y:A. B y). C x) = (UN y:A. UN x: B y. C x)"; |
469 |
by (Blast_tac 1); |
|
470 |
qed "UN_UN_flatten"; |
|
471 |
||
| 4059 | 472 |
goal thy "(INT x:insert a A. B x) = B a Int INTER A B"; |
| 2891 | 473 |
by (Blast_tac 1); |
| 1531 | 474 |
qed "INT_insert"; |
475 |
Addsimps[INT_insert]; |
|
|
1179
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
nipkow
parents:
923
diff
changeset
|
476 |
|
| 4059 | 477 |
goal thy |
| 2021 | 478 |
"!!A. A~={} ==> (INT x:A. insert a (B x)) = insert a (INT x:A. B x)";
|
| 2891 | 479 |
by (Blast_tac 1); |
| 2021 | 480 |
qed "INT_insert_distrib"; |
481 |
||
| 4059 | 482 |
goal thy "Union(B``A) = (UN x:A. B(x))"; |
| 2891 | 483 |
by (Blast_tac 1); |
| 923 | 484 |
qed "Union_image_eq"; |
485 |
||
| 4059 | 486 |
goal thy "Inter(B``A) = (INT x:A. B(x))"; |
| 2891 | 487 |
by (Blast_tac 1); |
| 923 | 488 |
qed "Inter_image_eq"; |
489 |
||
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
490 |
goal thy "!!A. A~={} ==> (UN y:A. c) = c";
|
| 2891 | 491 |
by (Blast_tac 1); |
| 923 | 492 |
qed "UN_constant"; |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
493 |
Addsimps[UN_constant]; |
| 923 | 494 |
|
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
495 |
goal thy "!!A. A~={} ==> (INT y:A. c) = c";
|
| 2891 | 496 |
by (Blast_tac 1); |
| 923 | 497 |
qed "INT_constant"; |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
498 |
Addsimps[INT_constant]; |
| 923 | 499 |
|
| 4059 | 500 |
goal thy "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
|
| 2891 | 501 |
by (Blast_tac 1); |
| 923 | 502 |
qed "UN_eq"; |
503 |
||
504 |
(*Look: it has an EXISTENTIAL quantifier*) |
|
| 4059 | 505 |
goal thy "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
|
| 2891 | 506 |
by (Blast_tac 1); |
| 923 | 507 |
qed "INT_eq"; |
508 |
||
| 4059 | 509 |
goalw thy [o_def] "UNION A (g o f) = UNION (f``A) g"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
510 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
511 |
qed "UNION_o"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
512 |
|
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
513 |
|
| 923 | 514 |
(*Distributive laws...*) |
515 |
||
| 4059 | 516 |
goal thy "A Int Union(B) = (UN C:B. A Int C)"; |
| 2891 | 517 |
by (Blast_tac 1); |
| 923 | 518 |
qed "Int_Union"; |
519 |
||
| 4674 | 520 |
goal thy "Union(B) Int A = (UN C:B. C Int A)"; |
521 |
by (Blast_tac 1); |
|
522 |
qed "Int_Union2"; |
|
523 |
||
|
4306
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
paulson
parents:
4231
diff
changeset
|
524 |
(* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: |
| 923 | 525 |
Union of a family of unions **) |
| 4059 | 526 |
goal thy "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; |
| 2891 | 527 |
by (Blast_tac 1); |
| 923 | 528 |
qed "Un_Union_image"; |
529 |
||
530 |
(*Equivalent version*) |
|
| 4059 | 531 |
goal thy "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; |
| 2891 | 532 |
by (Blast_tac 1); |
| 923 | 533 |
qed "UN_Un_distrib"; |
534 |
||
| 4059 | 535 |
goal thy "A Un Inter(B) = (INT C:B. A Un C)"; |
| 2891 | 536 |
by (Blast_tac 1); |
| 923 | 537 |
qed "Un_Inter"; |
538 |
||
| 4059 | 539 |
goal thy "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; |
| 2891 | 540 |
by (Blast_tac 1); |
| 923 | 541 |
qed "Int_Inter_image"; |
542 |
||
543 |
(*Equivalent version*) |
|
| 4059 | 544 |
goal thy "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; |
| 2891 | 545 |
by (Blast_tac 1); |
| 923 | 546 |
qed "INT_Int_distrib"; |
547 |
||
548 |
(*Halmos, Naive Set Theory, page 35.*) |
|
| 4059 | 549 |
goal thy "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; |
| 2891 | 550 |
by (Blast_tac 1); |
| 923 | 551 |
qed "Int_UN_distrib"; |
552 |
||
| 4059 | 553 |
goal thy "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; |
| 2891 | 554 |
by (Blast_tac 1); |
| 923 | 555 |
qed "Un_INT_distrib"; |
556 |
||
| 4059 | 557 |
goal thy |
| 923 | 558 |
"(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; |
| 2891 | 559 |
by (Blast_tac 1); |
| 923 | 560 |
qed "Int_UN_distrib2"; |
561 |
||
| 4059 | 562 |
goal thy |
| 923 | 563 |
"(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; |
| 2891 | 564 |
by (Blast_tac 1); |
| 923 | 565 |
qed "Un_INT_distrib2"; |
566 |
||
| 2512 | 567 |
|
568 |
section"Bounded quantifiers"; |
|
569 |
||
| 3860 | 570 |
(** The following are not added to the default simpset because |
571 |
(a) they duplicate the body and (b) there are no similar rules for Int. **) |
|
| 2512 | 572 |
|
| 4059 | 573 |
goal thy "(ALL x:A Un B. P x) = ((ALL x:A. P x) & (ALL x:B. P x))"; |
| 2891 | 574 |
by (Blast_tac 1); |
| 2519 | 575 |
qed "ball_Un"; |
576 |
||
| 4059 | 577 |
goal thy "(EX x:A Un B. P x) = ((EX x:A. P x) | (EX x:B. P x))"; |
| 2891 | 578 |
by (Blast_tac 1); |
| 2519 | 579 |
qed "bex_Un"; |
| 2512 | 580 |
|
| 4771 | 581 |
goal thy "(ALL z: UNION A B. P z) = (ALL x:A. ALL z:B x. P z)"; |
582 |
by (Blast_tac 1); |
|
583 |
qed "ball_UN"; |
|
584 |
||
585 |
goal thy "(EX z: UNION A B. P z) = (EX x:A. EX z:B x. P z)"; |
|
586 |
by (Blast_tac 1); |
|
587 |
qed "bex_UN"; |
|
588 |
||
| 2512 | 589 |
|
| 1548 | 590 |
section "-"; |
| 923 | 591 |
|
| 4609 | 592 |
goal thy "A-B = A Int Compl B"; |
593 |
by (Blast_tac 1); |
|
| 4662 | 594 |
qed "Diff_eq"; |
| 4609 | 595 |
|
| 4059 | 596 |
goal thy "A-A = {}";
|
| 2891 | 597 |
by (Blast_tac 1); |
| 923 | 598 |
qed "Diff_cancel"; |
| 1531 | 599 |
Addsimps[Diff_cancel]; |
| 923 | 600 |
|
| 4674 | 601 |
goal thy "!!A B. A Int B = {} ==> A-B = A";
|
602 |
by (blast_tac (claset() addEs [equalityE]) 1); |
|
603 |
qed "Diff_triv"; |
|
604 |
||
| 4059 | 605 |
goal thy "{}-A = {}";
|
| 2891 | 606 |
by (Blast_tac 1); |
| 923 | 607 |
qed "empty_Diff"; |
| 1531 | 608 |
Addsimps[empty_Diff]; |
| 923 | 609 |
|
| 4059 | 610 |
goal thy "A-{} = A";
|
| 2891 | 611 |
by (Blast_tac 1); |
| 923 | 612 |
qed "Diff_empty"; |
| 1531 | 613 |
Addsimps[Diff_empty]; |
614 |
||
| 4059 | 615 |
goal thy "A-UNIV = {}";
|
| 2891 | 616 |
by (Blast_tac 1); |
| 1531 | 617 |
qed "Diff_UNIV"; |
618 |
Addsimps[Diff_UNIV]; |
|
619 |
||
| 4059 | 620 |
goal thy "!!x. x~:A ==> A - insert x B = A-B"; |
| 2891 | 621 |
by (Blast_tac 1); |
| 1531 | 622 |
qed "Diff_insert0"; |
623 |
Addsimps [Diff_insert0]; |
|
| 923 | 624 |
|
625 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
|
|
| 4059 | 626 |
goal thy "A - insert a B = A - B - {a}";
|
| 2891 | 627 |
by (Blast_tac 1); |
| 923 | 628 |
qed "Diff_insert"; |
629 |
||
630 |
(*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
|
|
| 4059 | 631 |
goal thy "A - insert a B = A - {a} - B";
|
| 2891 | 632 |
by (Blast_tac 1); |
| 923 | 633 |
qed "Diff_insert2"; |
634 |
||
| 4059 | 635 |
goal thy "insert x A - B = (if x:B then A-B else insert x (A-B))"; |
| 4686 | 636 |
by (Simp_tac 1); |
| 2891 | 637 |
by (Blast_tac 1); |
| 1531 | 638 |
qed "insert_Diff_if"; |
639 |
||
| 4059 | 640 |
goal thy "!!x. x:B ==> insert x A - B = A-B"; |
| 2891 | 641 |
by (Blast_tac 1); |
| 1531 | 642 |
qed "insert_Diff1"; |
643 |
Addsimps [insert_Diff1]; |
|
644 |
||
| 4059 | 645 |
goal thy "!!a. a:A ==> insert a (A-{a}) = A";
|
| 2922 | 646 |
by (Blast_tac 1); |
| 923 | 647 |
qed "insert_Diff"; |
648 |
||
| 4059 | 649 |
goal thy "A Int (B-A) = {}";
|
| 2891 | 650 |
by (Blast_tac 1); |
| 923 | 651 |
qed "Diff_disjoint"; |
| 1531 | 652 |
Addsimps[Diff_disjoint]; |
| 923 | 653 |
|
| 4059 | 654 |
goal thy "!!A. A<=B ==> A Un (B-A) = B"; |
| 2891 | 655 |
by (Blast_tac 1); |
| 923 | 656 |
qed "Diff_partition"; |
657 |
||
| 4059 | 658 |
goal thy "!!A. [| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; |
| 2891 | 659 |
by (Blast_tac 1); |
| 923 | 660 |
qed "double_diff"; |
661 |
||
| 4645 | 662 |
goal thy "A Un (B-A) = A Un B"; |
663 |
by (Blast_tac 1); |
|
664 |
qed "Un_Diff_cancel"; |
|
665 |
||
666 |
goal thy "(B-A) Un A = B Un A"; |
|
667 |
by (Blast_tac 1); |
|
668 |
qed "Un_Diff_cancel2"; |
|
669 |
||
670 |
Addsimps [Un_Diff_cancel, Un_Diff_cancel2]; |
|
671 |
||
| 4059 | 672 |
goal thy "A - (B Un C) = (A-B) Int (A-C)"; |
| 2891 | 673 |
by (Blast_tac 1); |
| 923 | 674 |
qed "Diff_Un"; |
675 |
||
| 4059 | 676 |
goal thy "A - (B Int C) = (A-B) Un (A-C)"; |
| 2891 | 677 |
by (Blast_tac 1); |
| 923 | 678 |
qed "Diff_Int"; |
679 |
||
| 4059 | 680 |
goal thy "(A Un B) - C = (A - C) Un (B - C)"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
681 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
682 |
qed "Un_Diff"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
683 |
|
| 4662 | 684 |
goal thy "(A Int B) - C = A Int (B - C)"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
685 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
686 |
qed "Int_Diff"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
687 |
|
| 4748 | 688 |
goal thy "C Int (A-B) = (C Int A) - (C Int B)"; |
689 |
by (Blast_tac 1); |
|
690 |
qed "Diff_Int_distrib"; |
|
691 |
||
| 4645 | 692 |
goal thy "(A-B) Int C = (A Int C) - (B Int C)"; |
693 |
by (Blast_tac 1); |
|
| 4748 | 694 |
qed "Diff_Int_distrib2"; |
| 4645 | 695 |
|
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
696 |
|
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
697 |
section "Miscellany"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
698 |
|
| 4059 | 699 |
goal thy "(A = B) = ((A <= (B::'a set)) & (B<=A))"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
700 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
701 |
qed "set_eq_subset"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
702 |
|
| 4059 | 703 |
goal thy "A <= B = (! t. t:A --> t:B)"; |
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
704 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
705 |
qed "subset_iff"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
706 |
|
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
707 |
goalw thy [psubset_def] "((A::'a set) <= B) = ((A < B) | (A=B))"; |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
708 |
by (Blast_tac 1); |
|
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
2922
diff
changeset
|
709 |
qed "subset_iff_psubset_eq"; |
| 2021 | 710 |
|
| 4059 | 711 |
goal thy "(!x. x ~: A) = (A={})";
|
| 4423 | 712 |
by (Blast_tac 1); |
|
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3860
diff
changeset
|
713 |
qed "all_not_in_conv"; |
| 3907 | 714 |
AddIffs [all_not_in_conv]; |
|
3896
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
nipkow
parents:
3860
diff
changeset
|
715 |
|
| 4059 | 716 |
goalw thy [Pow_def] "Pow {} = {{}}";
|
|
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4423
diff
changeset
|
717 |
by Auto_tac; |
|
3348
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
718 |
qed "Pow_empty"; |
|
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
719 |
Addsimps [Pow_empty]; |
|
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
720 |
|
| 4059 | 721 |
goal thy "Pow (insert a A) = Pow A Un (insert a `` Pow A)"; |
| 3724 | 722 |
by Safe_tac; |
| 3457 | 723 |
by (etac swap 1); |
|
3348
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
724 |
by (res_inst_tac [("x", "x-{a}")] image_eqI 1);
|
|
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
725 |
by (ALLGOALS Blast_tac); |
|
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
726 |
qed "Pow_insert"; |
|
3f9a806f061e
Two useful facts about Powersets suggested by Florian Kammueller
paulson
parents:
3222
diff
changeset
|
727 |
|
| 2021 | 728 |
|
729 |
(** Miniscoping: pushing in big Unions and Intersections **) |
|
730 |
local |
|
| 4059 | 731 |
fun prover s = prove_goal thy s (fn _ => [Blast_tac 1]) |
| 2021 | 732 |
in |
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
733 |
val UN_simps = map prover |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
734 |
["!!C. C ~= {} ==> (UN x:C. insert a (B x)) = insert a (UN x:C. B x)",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
735 |
"!!C. C ~= {} ==> (UN x:C. A x Un B) = ((UN x:C. A x) Un B)",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
736 |
"!!C. C ~= {} ==> (UN x:C. A Un B x) = (A Un (UN x:C. B x))",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
737 |
"(UN x:C. A x Int B) = ((UN x:C. A x) Int B)", |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
738 |
"(UN x:C. A Int B x) = (A Int (UN x:C. B x))", |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
739 |
"(UN x:C. A x - B) = ((UN x:C. A x) - B)", |
| 4231 | 740 |
"(UN x:C. A - B x) = (A - (INT x:C. B x))", |
741 |
"(UN x:f``A. B x) = (UN a:A. B(f a))"]; |
|
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
742 |
|
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
743 |
val INT_simps = map prover |
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
744 |
["!!C. C ~= {} ==> (INT x:C. A x Int B) = ((INT x:C. A x) Int B)",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
745 |
"!!C. C ~= {} ==> (INT x:C. A Int B x) = (A Int (INT x:C. B x))",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
746 |
"!!C. C ~= {} ==> (INT x:C. A x - B) = ((INT x:C. A x) - B)",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
747 |
"!!C. C ~= {} ==> (INT x:C. A - B x) = (A - (UN x:C. B x))",
|
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
748 |
"(INT x:C. insert a (B x)) = insert a (INT x:C. B x)", |
|
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
749 |
"(INT x:C. A x Un B) = ((INT x:C. A x) Un B)", |
| 4231 | 750 |
"(INT x:C. A Un B x) = (A Un (INT x:C. B x))", |
751 |
"(INT x:f``A. B x) = (INT a:A. B(f a))"]; |
|
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
752 |
|
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
753 |
|
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
754 |
val ball_simps = map prover |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
755 |
["(ALL x:A. P x | Q) = ((ALL x:A. P x) | Q)", |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
756 |
"(ALL x:A. P | Q x) = (P | (ALL x:A. Q x))", |
| 3422 | 757 |
"(ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))", |
758 |
"(ALL x:A. P x --> Q) = ((EX x:A. P x) --> Q)", |
|
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
759 |
"(ALL x:{}. P x) = True",
|
| 4136 | 760 |
"(ALL x:UNIV. P x) = (ALL x. P x)", |
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
761 |
"(ALL x:insert a B. P x) = (P(a) & (ALL x:B. P x))", |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
762 |
"(ALL x:Union(A). P x) = (ALL y:A. ALL x:y. P x)", |
| 3860 | 763 |
"(ALL x:Collect Q. P x) = (ALL x. Q x --> P x)", |
764 |
"(ALL x:f``A. P x) = (ALL x:A. P(f x))", |
|
765 |
"(~(ALL x:A. P x)) = (EX x:A. ~P x)"]; |
|
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
766 |
|
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
767 |
val ball_conj_distrib = |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
768 |
prover "(ALL x:A. P x & Q x) = ((ALL x:A. P x) & (ALL x:A. Q x))"; |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
769 |
|
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
770 |
val bex_simps = map prover |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
771 |
["(EX x:A. P x & Q) = ((EX x:A. P x) & Q)", |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
772 |
"(EX x:A. P & Q x) = (P & (EX x:A. Q x))", |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
773 |
"(EX x:{}. P x) = False",
|
| 4136 | 774 |
"(EX x:UNIV. P x) = (EX x. P x)", |
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
775 |
"(EX x:insert a B. P x) = (P(a) | (EX x:B. P x))", |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
776 |
"(EX x:Union(A). P x) = (EX y:A. EX x:y. P x)", |
| 3860 | 777 |
"(EX x:Collect Q. P x) = (EX x. Q x & P x)", |
778 |
"(EX x:f``A. P x) = (EX x:A. P(f x))", |
|
779 |
"(~(EX x:A. P x)) = (ALL x:A. ~P x)"]; |
|
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
780 |
|
| 3426 | 781 |
val bex_disj_distrib = |
|
2513
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
782 |
prover "(EX x:A. P x | Q x) = ((EX x:A. P x) | (EX x:A. Q x))"; |
|
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
paulson
parents:
2512
diff
changeset
|
783 |
|
| 2021 | 784 |
end; |
785 |
||
|
4159
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
paulson
parents:
4136
diff
changeset
|
786 |
Addsimps (UN_simps @ INT_simps @ ball_simps @ bex_simps); |