| author | wenzelm | 
| Wed, 31 May 2000 14:30:28 +0200 | |
| changeset 9011 | 0cfc347f8d19 | 
| parent 8993 | cbfebff56cc0 | 
| child 9077 | 5bf9b0d6df8a | 
| 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: 
1748diff
changeset | 11 | AddSIs [equalityI]; | 
| 
852093aeb0ab
Replaced fast_tac by Fast_tac (which uses default claset)
 berghofe parents: 
1748diff
changeset | 12 | |
| 1548 | 13 | section "{}";
 | 
| 14 | ||
| 7824 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 15 | (*supersedes Collect_False_empty*) | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 16 | Goal "{s. P} = (if P then UNIV else {})";
 | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 17 | by (Force_tac 1); | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 18 | qed "Collect_const"; | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 19 | Addsimps [Collect_const]; | 
| 1531 | 20 | |
| 5069 | 21 | Goal "(A <= {}) = (A = {})";
 | 
| 2891 | 22 | by (Blast_tac 1); | 
| 1531 | 23 | qed "subset_empty"; | 
| 24 | Addsimps [subset_empty]; | |
| 25 | ||
| 5069 | 26 | Goalw [psubset_def] "~ (A < {})";
 | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 27 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 28 | qed "not_psubset_empty"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 29 | AddIffs [not_psubset_empty]; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 30 | |
| 7958 | 31 | Goal "(Collect P = {}) = (!x. ~ P x)";
 | 
| 32 | by Auto_tac; | |
| 33 | qed "Collect_empty_eq"; | |
| 34 | Addsimps[Collect_empty_eq]; | |
| 35 | ||
| 8993 | 36 | Goal "{x. ~ (P x)} = - {x. P x}";
 | 
| 37 | by (Blast_tac 1); | |
| 38 | qed "Collect_neg_eq"; | |
| 39 | ||
| 5069 | 40 | Goal "{x. P x | Q x} = {x. P x} Un {x. Q x}";
 | 
| 4748 | 41 | by (Blast_tac 1); | 
| 42 | qed "Collect_disj_eq"; | |
| 43 | ||
| 5069 | 44 | Goal "{x. P x & Q x} = {x. P x} Int {x. Q x}";
 | 
| 4748 | 45 | by (Blast_tac 1); | 
| 46 | qed "Collect_conj_eq"; | |
| 47 | ||
| 7845 | 48 | Goal "{x. ALL y. P x y} = (INT y. {x. P x y})";
 | 
| 49 | by (Blast_tac 1); | |
| 50 | qed "Collect_all_eq"; | |
| 51 | ||
| 52 | Goal "{x. ALL y: A. P x y} = (INT y: A. {x. P x y})";
 | |
| 53 | by (Blast_tac 1); | |
| 54 | qed "Collect_ball_eq"; | |
| 55 | ||
| 56 | Goal "{x. EX y. P x y} = (UN y. {x. P x y})";
 | |
| 57 | by (Blast_tac 1); | |
| 58 | qed "Collect_ex_eq"; | |
| 59 | ||
| 60 | Goal "{x. EX y: A. P x y} = (UN y: A. {x. P x y})";
 | |
| 61 | by (Blast_tac 1); | |
| 62 | qed "Collect_bex_eq"; | |
| 63 | ||
| 4748 | 64 | |
| 1548 | 65 | section "insert"; | 
| 923 | 66 | |
| 1531 | 67 | (*NOT SUITABLE FOR REWRITING since {a} == insert a {}*)
 | 
| 5069 | 68 | Goal "insert a A = {a} Un A";
 | 
| 2891 | 69 | by (Blast_tac 1); | 
| 1531 | 70 | qed "insert_is_Un"; | 
| 71 | ||
| 5069 | 72 | Goal "insert a A ~= {}";
 | 
| 4089 | 73 | by (blast_tac (claset() addEs [equalityCE]) 1); | 
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 74 | qed"insert_not_empty"; | 
| 1531 | 75 | Addsimps[insert_not_empty]; | 
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 76 | |
| 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 77 | bind_thm("empty_not_insert",insert_not_empty RS not_sym);
 | 
| 1531 | 78 | Addsimps[empty_not_insert]; | 
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 79 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 80 | Goal "a:A ==> insert a A = A"; | 
| 2891 | 81 | by (Blast_tac 1); | 
| 923 | 82 | qed "insert_absorb"; | 
| 6832 | 83 | (* Addsimps [insert_absorb] causes recursive calls | 
| 84 | when there are nested inserts, with QUADRATIC running time | |
| 4605 | 85 | *) | 
| 923 | 86 | |
| 5069 | 87 | Goal "insert x (insert x A) = insert x A"; | 
| 2891 | 88 | by (Blast_tac 1); | 
| 1531 | 89 | qed "insert_absorb2"; | 
| 90 | Addsimps [insert_absorb2]; | |
| 91 | ||
| 5069 | 92 | Goal "insert x (insert y A) = insert y (insert x A)"; | 
| 2891 | 93 | by (Blast_tac 1); | 
| 1879 | 94 | qed "insert_commute"; | 
| 95 | ||
| 5069 | 96 | Goal "(insert x A <= B) = (x:B & A <= B)"; | 
| 2891 | 97 | by (Blast_tac 1); | 
| 923 | 98 | qed "insert_subset"; | 
| 1531 | 99 | Addsimps[insert_subset]; | 
| 100 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 101 | Goal "insert a A ~= insert a B ==> A ~= B"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 102 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 103 | qed "insert_lim"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 104 | |
| 1531 | 105 | (* use new B rather than (A-{a}) to avoid infinite unfolding *)
 | 
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 106 | Goal "a:A ==> ? B. A = insert a B & a ~: B"; | 
| 1553 | 107 | by (res_inst_tac [("x","A-{a}")] exI 1);
 | 
| 2891 | 108 | by (Blast_tac 1); | 
| 1531 | 109 | qed "mk_disjoint_insert"; | 
| 923 | 110 | |
| 4882 | 111 | bind_thm ("insert_Collect", prove_goal thy 
 | 
| 5590 | 112 | 	  "insert a (Collect P) = {u. u ~= a --> P u}" (K [Auto_tac]));
 | 
| 4882 | 113 | |
| 5941 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 114 | Goal "u: A ==> (UN x:A. insert a (B x)) = insert a (UN x:A. B x)"; | 
| 2891 | 115 | by (Blast_tac 1); | 
| 1843 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 paulson parents: 
1786diff
changeset | 116 | qed "UN_insert_distrib"; | 
| 
a6d7aef48c2f
Removed the unused eq_cs, and added some distributive laws
 paulson parents: 
1786diff
changeset | 117 | |
| 1660 | 118 | section "``"; | 
| 923 | 119 | |
| 5069 | 120 | Goal "f``{} = {}";
 | 
| 2891 | 121 | by (Blast_tac 1); | 
| 923 | 122 | qed "image_empty"; | 
| 1531 | 123 | Addsimps[image_empty]; | 
| 923 | 124 | |
| 5069 | 125 | Goal "f``insert a B = insert (f a) (f``B)"; | 
| 2891 | 126 | by (Blast_tac 1); | 
| 923 | 127 | qed "image_insert"; | 
| 1531 | 128 | Addsimps[image_insert]; | 
| 923 | 129 | |
| 6292 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 130 | Goal "x:A ==> (%x. c) `` A = {c}";
 | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 131 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 132 | qed "image_constant"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 133 | |
| 5069 | 134 | Goal "f``(g``A) = (%x. f (g x)) `` A"; | 
| 3457 | 135 | by (Blast_tac 1); | 
| 4059 | 136 | qed "image_image"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 137 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 138 | Goal "x:A ==> insert (f x) (f``A) = f``A"; | 
| 2891 | 139 | by (Blast_tac 1); | 
| 1884 | 140 | qed "insert_image"; | 
| 141 | Addsimps [insert_image]; | |
| 142 | ||
| 5069 | 143 | Goal "(f``A = {}) = (A = {})";
 | 
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 144 | by (blast_tac (claset() addSEs [equalityCE]) 1); | 
| 3415 
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
 nipkow parents: 
3384diff
changeset | 145 | qed "image_is_empty"; | 
| 
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
 nipkow parents: 
3384diff
changeset | 146 | AddIffs [image_is_empty]; | 
| 
c068bd2f0bbd
Moved image_is_empty from Finite.ML to equalities.ML
 nipkow parents: 
3384diff
changeset | 147 | |
| 5281 | 148 | Goal "f `` {x. P x} = {f x | x. P x}";
 | 
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 149 | by (Blast_tac 1); | 
| 5281 | 150 | qed "image_Collect"; | 
| 151 | Addsimps [image_Collect]; | |
| 152 | ||
| 5590 | 153 | Goalw [image_def] "(%x. if P x then f x else g x) `` S \ | 
| 154 | \                = (f `` (S Int {x. P x})) Un (g `` (S Int {x. ~(P x)}))";
 | |
| 4686 | 155 | by (Simp_tac 1); | 
| 2891 | 156 | by (Blast_tac 1); | 
| 1748 | 157 | qed "if_image_distrib"; | 
| 158 | Addsimps[if_image_distrib]; | |
| 159 | ||
| 5590 | 160 | val prems = Goal "[|M = N; !!x. x:N ==> f x = g x|] ==> f``M = g``N"; | 
| 4136 | 161 | by (simp_tac (simpset() addsimps image_def::prems) 1); | 
| 162 | qed "image_cong"; | |
| 163 | ||
| 1748 | 164 | |
| 7958 | 165 | section "range"; | 
| 166 | ||
| 167 | Goal "{u. ? x. u = f x} = range f";
 | |
| 168 | by Auto_tac; | |
| 169 | qed "full_SetCompr_eq"; | |
| 170 | ||
| 8161 | 171 | Goal "range (%x. f (g x)) = f``range g"; | 
| 172 | by (stac image_image 1); | |
| 173 | by (Simp_tac 1); | |
| 174 | qed "range_composition"; | |
| 175 | Addsimps[range_composition]; | |
| 7958 | 176 | |
| 1548 | 177 | section "Int"; | 
| 923 | 178 | |
| 5069 | 179 | Goal "A Int A = A"; | 
| 2891 | 180 | by (Blast_tac 1); | 
| 923 | 181 | qed "Int_absorb"; | 
| 1531 | 182 | Addsimps[Int_absorb]; | 
| 923 | 183 | |
| 5590 | 184 | Goal "A Int (A Int B) = A Int B"; | 
| 4609 | 185 | by (Blast_tac 1); | 
| 186 | qed "Int_left_absorb"; | |
| 187 | ||
| 5590 | 188 | Goal "A Int B = B Int A"; | 
| 2891 | 189 | by (Blast_tac 1); | 
| 923 | 190 | qed "Int_commute"; | 
| 191 | ||
| 5069 | 192 | Goal "A Int (B Int C) = B Int (A Int C)"; | 
| 4609 | 193 | by (Blast_tac 1); | 
| 194 | qed "Int_left_commute"; | |
| 195 | ||
| 5590 | 196 | Goal "(A Int B) Int C = A Int (B Int C)"; | 
| 2891 | 197 | by (Blast_tac 1); | 
| 923 | 198 | qed "Int_assoc"; | 
| 199 | ||
| 4609 | 200 | (*Intersection is an AC-operator*) | 
| 7648 | 201 | bind_thms ("Int_ac", [Int_assoc, Int_left_absorb, Int_commute, Int_left_commute]);
 | 
| 4609 | 202 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 203 | Goal "B<=A ==> A Int B = B"; | 
| 4662 | 204 | by (Blast_tac 1); | 
| 205 | qed "Int_absorb1"; | |
| 206 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 207 | Goal "A<=B ==> A Int B = A"; | 
| 4662 | 208 | by (Blast_tac 1); | 
| 209 | qed "Int_absorb2"; | |
| 210 | ||
| 5069 | 211 | Goal "{} Int B = {}";
 | 
| 2891 | 212 | by (Blast_tac 1); | 
| 923 | 213 | qed "Int_empty_left"; | 
| 1531 | 214 | Addsimps[Int_empty_left]; | 
| 923 | 215 | |
| 5069 | 216 | Goal "A Int {} = {}";
 | 
| 2891 | 217 | by (Blast_tac 1); | 
| 923 | 218 | qed "Int_empty_right"; | 
| 1531 | 219 | Addsimps[Int_empty_right]; | 
| 220 | ||
| 5490 | 221 | Goal "(A Int B = {}) = (A <= -B)";
 | 
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 222 | by (blast_tac (claset() addSEs [equalityCE]) 1); | 
| 3356 | 223 | qed "disjoint_eq_subset_Compl"; | 
| 224 | ||
| 7877 | 225 | Goal "(A Int B = {}) = (ALL x:A. ALL y:B. x ~= y)";
 | 
| 226 | by (Blast_tac 1); | |
| 227 | qed "disjoint_iff_not_equal"; | |
| 228 | ||
| 5069 | 229 | Goal "UNIV Int B = B"; | 
| 2891 | 230 | by (Blast_tac 1); | 
| 1531 | 231 | qed "Int_UNIV_left"; | 
| 232 | Addsimps[Int_UNIV_left]; | |
| 233 | ||
| 5069 | 234 | Goal "A Int UNIV = A"; | 
| 2891 | 235 | by (Blast_tac 1); | 
| 1531 | 236 | qed "Int_UNIV_right"; | 
| 237 | Addsimps[Int_UNIV_right]; | |
| 923 | 238 | |
| 5069 | 239 | Goal "A Int B = Inter{A,B}";
 | 
| 4634 | 240 | by (Blast_tac 1); | 
| 241 | qed "Int_eq_Inter"; | |
| 242 | ||
| 5590 | 243 | Goal "A Int (B Un C) = (A Int B) Un (A Int C)"; | 
| 2891 | 244 | by (Blast_tac 1); | 
| 923 | 245 | qed "Int_Un_distrib"; | 
| 246 | ||
| 5590 | 247 | Goal "(B Un C) Int A = (B Int A) Un (C Int A)"; | 
| 2891 | 248 | by (Blast_tac 1); | 
| 1618 | 249 | qed "Int_Un_distrib2"; | 
| 250 | ||
| 5069 | 251 | Goal "(A Int B = UNIV) = (A = UNIV & B = UNIV)"; | 
| 4089 | 252 | by (blast_tac (claset() addEs [equalityCE]) 1); | 
| 1531 | 253 | qed "Int_UNIV"; | 
| 254 | Addsimps[Int_UNIV]; | |
| 255 | ||
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 256 | Goal "(C <= A Int B) = (C <= A & C <= B)"; | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 257 | by (Blast_tac 1); | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 258 | qed "Int_subset_iff"; | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 259 | |
| 7713 | 260 | Goal "(x : A Int {x. P x}) = (x : A & P x)";
 | 
| 261 | by (Blast_tac 1); | |
| 262 | qed "Int_Collect"; | |
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 263 | |
| 1548 | 264 | section "Un"; | 
| 923 | 265 | |
| 5069 | 266 | Goal "A Un A = A"; | 
| 2891 | 267 | by (Blast_tac 1); | 
| 923 | 268 | qed "Un_absorb"; | 
| 1531 | 269 | Addsimps[Un_absorb]; | 
| 923 | 270 | |
| 5069 | 271 | Goal " A Un (A Un B) = A Un B"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 272 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 273 | qed "Un_left_absorb"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 274 | |
| 5590 | 275 | Goal "A Un B = B Un A"; | 
| 2891 | 276 | by (Blast_tac 1); | 
| 923 | 277 | qed "Un_commute"; | 
| 278 | ||
| 5069 | 279 | Goal "A Un (B Un C) = B Un (A Un C)"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 280 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 281 | qed "Un_left_commute"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 282 | |
| 5590 | 283 | Goal "(A Un B) Un C = A Un (B Un C)"; | 
| 2891 | 284 | by (Blast_tac 1); | 
| 923 | 285 | qed "Un_assoc"; | 
| 286 | ||
| 4609 | 287 | (*Union is an AC-operator*) | 
| 7648 | 288 | bind_thms ("Un_ac", [Un_assoc, Un_left_absorb, Un_commute, Un_left_commute]);
 | 
| 4609 | 289 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 290 | Goal "A<=B ==> A Un B = B"; | 
| 4662 | 291 | by (Blast_tac 1); | 
| 292 | qed "Un_absorb1"; | |
| 293 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 294 | Goal "B<=A ==> A Un B = A"; | 
| 4662 | 295 | by (Blast_tac 1); | 
| 296 | qed "Un_absorb2"; | |
| 297 | ||
| 5069 | 298 | Goal "{} Un B = B";
 | 
| 2891 | 299 | by (Blast_tac 1); | 
| 923 | 300 | qed "Un_empty_left"; | 
| 1531 | 301 | Addsimps[Un_empty_left]; | 
| 923 | 302 | |
| 5069 | 303 | Goal "A Un {} = A";
 | 
| 2891 | 304 | by (Blast_tac 1); | 
| 923 | 305 | qed "Un_empty_right"; | 
| 1531 | 306 | Addsimps[Un_empty_right]; | 
| 307 | ||
| 5069 | 308 | Goal "UNIV Un B = UNIV"; | 
| 2891 | 309 | by (Blast_tac 1); | 
| 1531 | 310 | qed "Un_UNIV_left"; | 
| 311 | Addsimps[Un_UNIV_left]; | |
| 312 | ||
| 5069 | 313 | Goal "A Un UNIV = UNIV"; | 
| 2891 | 314 | by (Blast_tac 1); | 
| 1531 | 315 | qed "Un_UNIV_right"; | 
| 316 | Addsimps[Un_UNIV_right]; | |
| 923 | 317 | |
| 5069 | 318 | Goal "A Un B = Union{A,B}";
 | 
| 4634 | 319 | by (Blast_tac 1); | 
| 320 | qed "Un_eq_Union"; | |
| 321 | ||
| 5069 | 322 | Goal "(insert a B) Un C = insert a (B Un C)"; | 
| 2891 | 323 | by (Blast_tac 1); | 
| 923 | 324 | qed "Un_insert_left"; | 
| 3384 
5ef99c94e1fb
Now Un_insert_left, Un_insert_right are default rewrite rules
 paulson parents: 
3356diff
changeset | 325 | Addsimps[Un_insert_left]; | 
| 923 | 326 | |
| 5069 | 327 | Goal "A Un (insert a B) = insert a (A Un B)"; | 
| 2891 | 328 | by (Blast_tac 1); | 
| 1917 | 329 | qed "Un_insert_right"; | 
| 3384 
5ef99c94e1fb
Now Un_insert_left, Un_insert_right are default rewrite rules
 paulson parents: 
3356diff
changeset | 330 | Addsimps[Un_insert_right]; | 
| 1917 | 331 | |
| 5069 | 332 | Goal "(insert a B) Int C = (if a:C then insert a (B Int C) \ | 
| 5590 | 333 | \ else B Int C)"; | 
| 4686 | 334 | by (Simp_tac 1); | 
| 3356 | 335 | by (Blast_tac 1); | 
| 336 | qed "Int_insert_left"; | |
| 337 | ||
| 5069 | 338 | Goal "A Int (insert a B) = (if a:A then insert a (A Int B) \ | 
| 5590 | 339 | \ else A Int B)"; | 
| 4686 | 340 | by (Simp_tac 1); | 
| 3356 | 341 | by (Blast_tac 1); | 
| 342 | qed "Int_insert_right"; | |
| 343 | ||
| 5590 | 344 | Goal "A Un (B Int C) = (A Un B) Int (A Un C)"; | 
| 2891 | 345 | by (Blast_tac 1); | 
| 923 | 346 | qed "Un_Int_distrib"; | 
| 347 | ||
| 5590 | 348 | Goal "(B Int C) Un A = (B Un A) Int (C Un A)"; | 
| 4609 | 349 | by (Blast_tac 1); | 
| 350 | qed "Un_Int_distrib2"; | |
| 351 | ||
| 5590 | 352 | Goal "(A Int B) Un (B Int C) Un (C Int A) = \ | 
| 353 | \ (A Un B) Int (B Un C) Int (C Un A)"; | |
| 2891 | 354 | by (Blast_tac 1); | 
| 923 | 355 | qed "Un_Int_crazy"; | 
| 356 | ||
| 5069 | 357 | Goal "(A<=B) = (A Un B = B)"; | 
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 358 | by (blast_tac (claset() addSEs [equalityCE]) 1); | 
| 923 | 359 | qed "subset_Un_eq"; | 
| 360 | ||
| 5069 | 361 | Goal "(A <= insert b C) = (A <= C | b:A & A-{b} <= C)";
 | 
| 2891 | 362 | by (Blast_tac 1); | 
| 923 | 363 | qed "subset_insert_iff"; | 
| 364 | ||
| 5069 | 365 | Goal "(A Un B = {}) = (A = {} & B = {})";
 | 
| 4089 | 366 | by (blast_tac (claset() addEs [equalityCE]) 1); | 
| 923 | 367 | qed "Un_empty"; | 
| 1531 | 368 | Addsimps[Un_empty]; | 
| 923 | 369 | |
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 370 | Goal "(A Un B <= C) = (A <= C & B <= C)"; | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 371 | by (Blast_tac 1); | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 372 | qed "Un_subset_iff"; | 
| 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 373 | |
| 5331 | 374 | Goal "(A - B) Un (A Int B) = A"; | 
| 375 | by (Blast_tac 1); | |
| 376 | qed "Un_Diff_Int"; | |
| 377 | ||
| 5319 
7356d0c88b1b
Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
 paulson parents: 
5316diff
changeset | 378 | |
| 5931 | 379 | section "Set complement"; | 
| 923 | 380 | |
| 8333 | 381 | Goal "A Int -A = {}";
 | 
| 2891 | 382 | by (Blast_tac 1); | 
| 923 | 383 | qed "Compl_disjoint"; | 
| 8333 | 384 | |
| 385 | Goal "-A Int A = {}";
 | |
| 386 | by (Blast_tac 1); | |
| 387 | qed "Compl_disjoint2"; | |
| 388 | Addsimps[Compl_disjoint, Compl_disjoint2]; | |
| 923 | 389 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 390 | Goal "A Un (-A) = UNIV"; | 
| 2891 | 391 | by (Blast_tac 1); | 
| 923 | 392 | qed "Compl_partition"; | 
| 393 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 394 | Goal "- (-A) = (A:: 'a set)"; | 
| 2891 | 395 | by (Blast_tac 1); | 
| 923 | 396 | qed "double_complement"; | 
| 1531 | 397 | Addsimps[double_complement]; | 
| 923 | 398 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 399 | Goal "-(A Un B) = (-A) Int (-B)"; | 
| 2891 | 400 | by (Blast_tac 1); | 
| 923 | 401 | qed "Compl_Un"; | 
| 402 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 403 | Goal "-(A Int B) = (-A) Un (-B)"; | 
| 2891 | 404 | by (Blast_tac 1); | 
| 923 | 405 | qed "Compl_Int"; | 
| 406 | ||
| 5490 | 407 | Goal "-(UN x:A. B(x)) = (INT x:A. -B(x))"; | 
| 2891 | 408 | by (Blast_tac 1); | 
| 923 | 409 | qed "Compl_UN"; | 
| 410 | ||
| 5490 | 411 | Goal "-(INT x:A. B(x)) = (UN x:A. -B(x))"; | 
| 2891 | 412 | by (Blast_tac 1); | 
| 923 | 413 | qed "Compl_INT"; | 
| 414 | ||
| 4615 | 415 | Addsimps [Compl_Un, Compl_Int, Compl_UN, Compl_INT]; | 
| 416 | ||
| 8121 | 417 | Goal "(A <= -A) = (A = {})";
 | 
| 418 | by (Blast_tac 1); | |
| 419 | qed "subset_Compl_self_eq"; | |
| 420 | ||
| 923 | 421 | (*Halmos, Naive Set Theory, page 16.*) | 
| 5069 | 422 | Goal "((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: 
4231diff
changeset | 423 | by (blast_tac (claset() addSEs [equalityCE]) 1); | 
| 923 | 424 | qed "Un_Int_assoc_eq"; | 
| 425 | ||
| 8333 | 426 | Goal "-UNIV = {}";
 | 
| 427 | by (Blast_tac 1); | |
| 428 | qed "Compl_UNIV_eq"; | |
| 429 | ||
| 430 | Goal "-{} = UNIV";
 | |
| 431 | by (Blast_tac 1); | |
| 432 | qed "Compl_empty_eq"; | |
| 433 | ||
| 434 | Addsimps [Compl_UNIV_eq, Compl_empty_eq]; | |
| 435 | ||
| 923 | 436 | |
| 1548 | 437 | section "Union"; | 
| 923 | 438 | |
| 5069 | 439 | Goal "Union({}) = {}";
 | 
| 2891 | 440 | by (Blast_tac 1); | 
| 923 | 441 | qed "Union_empty"; | 
| 1531 | 442 | Addsimps[Union_empty]; | 
| 443 | ||
| 5069 | 444 | Goal "Union(UNIV) = UNIV"; | 
| 2891 | 445 | by (Blast_tac 1); | 
| 1531 | 446 | qed "Union_UNIV"; | 
| 447 | Addsimps[Union_UNIV]; | |
| 923 | 448 | |
| 5069 | 449 | Goal "Union(insert a B) = a Un Union(B)"; | 
| 2891 | 450 | by (Blast_tac 1); | 
| 923 | 451 | qed "Union_insert"; | 
| 1531 | 452 | Addsimps[Union_insert]; | 
| 923 | 453 | |
| 5069 | 454 | Goal "Union(A Un B) = Union(A) Un Union(B)"; | 
| 2891 | 455 | by (Blast_tac 1); | 
| 923 | 456 | qed "Union_Un_distrib"; | 
| 1531 | 457 | Addsimps[Union_Un_distrib]; | 
| 923 | 458 | |
| 5069 | 459 | Goal "Union(A Int B) <= Union(A) Int Union(B)"; | 
| 2891 | 460 | by (Blast_tac 1); | 
| 923 | 461 | qed "Union_Int_subset"; | 
| 462 | ||
| 5069 | 463 | Goal "(Union M = {}) = (! A : M. A = {})"; 
 | 
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 464 | by (blast_tac (claset() addEs [equalityCE]) 1); | 
| 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 465 | qed "Union_empty_conv"; | 
| 4003 | 466 | AddIffs [Union_empty_conv]; | 
| 467 | ||
| 5069 | 468 | Goal "(Union(C) Int A = {}) = (! B:C. B Int A = {})";
 | 
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 469 | by (blast_tac (claset() addSEs [equalityCE]) 1); | 
| 923 | 470 | qed "Union_disjoint"; | 
| 471 | ||
| 1548 | 472 | section "Inter"; | 
| 473 | ||
| 5069 | 474 | Goal "Inter({}) = UNIV";
 | 
| 2891 | 475 | by (Blast_tac 1); | 
| 1531 | 476 | qed "Inter_empty"; | 
| 477 | Addsimps[Inter_empty]; | |
| 478 | ||
| 5069 | 479 | Goal "Inter(UNIV) = {}";
 | 
| 2891 | 480 | by (Blast_tac 1); | 
| 1531 | 481 | qed "Inter_UNIV"; | 
| 482 | Addsimps[Inter_UNIV]; | |
| 483 | ||
| 5069 | 484 | Goal "Inter(insert a B) = a Int Inter(B)"; | 
| 2891 | 485 | by (Blast_tac 1); | 
| 1531 | 486 | qed "Inter_insert"; | 
| 487 | Addsimps[Inter_insert]; | |
| 488 | ||
| 5069 | 489 | Goal "Inter(A) Un Inter(B) <= Inter(A Int B)"; | 
| 2891 | 490 | by (Blast_tac 1); | 
| 1564 
822575c737bd
Deleted faulty comment; proved new rule Inter_Un_subset
 paulson parents: 
1553diff
changeset | 491 | qed "Inter_Un_subset"; | 
| 1531 | 492 | |
| 5069 | 493 | Goal "Inter(A Un B) = Inter(A) Int Inter(B)"; | 
| 2891 | 494 | by (Blast_tac 1); | 
| 923 | 495 | qed "Inter_Un_distrib"; | 
| 496 | ||
| 1548 | 497 | section "UN and INT"; | 
| 923 | 498 | |
| 499 | (*Basic identities*) | |
| 500 | ||
| 8030 | 501 | bind_thm ("not_empty", prove_goal thy "(A ~= {}) = (? x. x:A)" (K [Blast_tac 1]));
 | 
| 4136 | 502 | |
| 5069 | 503 | Goal "(UN x:{}. B x) = {}";
 | 
| 2891 | 504 | by (Blast_tac 1); | 
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 505 | qed "UN_empty"; | 
| 1531 | 506 | Addsimps[UN_empty]; | 
| 507 | ||
| 5069 | 508 | Goal "(UN x:A. {}) = {}";
 | 
| 3457 | 509 | by (Blast_tac 1); | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 510 | qed "UN_empty2"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 511 | Addsimps[UN_empty2]; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 512 | |
| 5069 | 513 | Goal "(UN x:A. {x}) = A";
 | 
| 4645 | 514 | by (Blast_tac 1); | 
| 515 | qed "UN_singleton"; | |
| 516 | Addsimps [UN_singleton]; | |
| 517 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 518 | Goal "k:I ==> A k Un (UN i:I. A i) = (UN i:I. A i)"; | 
| 4634 | 519 | by (Blast_tac 1); | 
| 520 | qed "UN_absorb"; | |
| 521 | ||
| 5069 | 522 | Goal "(INT x:{}. B x) = UNIV";
 | 
| 2891 | 523 | by (Blast_tac 1); | 
| 1531 | 524 | qed "INT_empty"; | 
| 525 | Addsimps[INT_empty]; | |
| 526 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 527 | Goal "k:I ==> A k Int (INT i:I. A i) = (INT i:I. A i)"; | 
| 4634 | 528 | by (Blast_tac 1); | 
| 529 | qed "INT_absorb"; | |
| 530 | ||
| 5069 | 531 | Goal "(UN x:insert a A. B x) = B a Un UNION A B"; | 
| 2891 | 532 | by (Blast_tac 1); | 
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 533 | qed "UN_insert"; | 
| 1531 | 534 | Addsimps[UN_insert]; | 
| 535 | ||
| 5999 | 536 | Goal "(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: 
2922diff
changeset | 537 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 538 | qed "UN_Un"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 539 | |
| 5069 | 540 | Goal "(UN x : (UN y:A. B y). C x) = (UN y:A. UN x: B y. C x)"; | 
| 4771 | 541 | by (Blast_tac 1); | 
| 542 | qed "UN_UN_flatten"; | |
| 543 | ||
| 6292 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 544 | Goal "((UN i:I. A i) <= B) = (ALL i:I. A i <= B)"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 545 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 546 | qed "UN_subset_iff"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 547 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 548 | Goal "(B <= (INT i:I. A i)) = (ALL i:I. B <= A i)"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 549 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 550 | qed "INT_subset_iff"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 551 | |
| 5069 | 552 | Goal "(INT x:insert a A. B x) = B a Int INTER A B"; | 
| 2891 | 553 | by (Blast_tac 1); | 
| 1531 | 554 | qed "INT_insert"; | 
| 555 | Addsimps[INT_insert]; | |
| 1179 
7678408f9751
Added insert_not_empty, UN_empty and UN_insert (to set_ss).
 nipkow parents: 
923diff
changeset | 556 | |
| 5999 | 557 | Goal "(INT i: A Un B. M i) = (INT i: A. M i) Int (INT i:B. M i)"; | 
| 558 | by (Blast_tac 1); | |
| 559 | qed "INT_Un"; | |
| 560 | ||
| 5941 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 561 | Goal "u: A ==> (INT x:A. insert a (B x)) = insert a (INT x:A. B x)"; | 
| 2891 | 562 | by (Blast_tac 1); | 
| 2021 | 563 | qed "INT_insert_distrib"; | 
| 564 | ||
| 5069 | 565 | Goal "Union(B``A) = (UN x:A. B(x))"; | 
| 2891 | 566 | by (Blast_tac 1); | 
| 923 | 567 | qed "Union_image_eq"; | 
| 6292 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 568 | Addsimps [Union_image_eq]; | 
| 923 | 569 | |
| 6283 | 570 | Goal "f `` Union S = (UN x:S. f `` x)"; | 
| 571 | by (Blast_tac 1); | |
| 8176 | 572 | qed "image_Union"; | 
| 6283 | 573 | |
| 5069 | 574 | Goal "Inter(B``A) = (INT x:A. B(x))"; | 
| 2891 | 575 | by (Blast_tac 1); | 
| 923 | 576 | qed "Inter_image_eq"; | 
| 6292 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 577 | Addsimps [Inter_image_eq]; | 
| 923 | 578 | |
| 8313 
c7a87e79e9b1
replaced UN_constant, INT_constant by unconditional versions that rewrite
 paulson parents: 
8176diff
changeset | 579 | Goal "(UN y:A. c) = (if A={} then {} else c)";
 | 
| 
c7a87e79e9b1
replaced UN_constant, INT_constant by unconditional versions that rewrite
 paulson parents: 
8176diff
changeset | 580 | by Auto_tac; | 
| 923 | 581 | qed "UN_constant"; | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4136diff
changeset | 582 | Addsimps[UN_constant]; | 
| 923 | 583 | |
| 8313 
c7a87e79e9b1
replaced UN_constant, INT_constant by unconditional versions that rewrite
 paulson parents: 
8176diff
changeset | 584 | Goal "(INT y:A. c) = (if A={} then UNIV else c)";
 | 
| 
c7a87e79e9b1
replaced UN_constant, INT_constant by unconditional versions that rewrite
 paulson parents: 
8176diff
changeset | 585 | by Auto_tac; | 
| 923 | 586 | qed "INT_constant"; | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4136diff
changeset | 587 | Addsimps[INT_constant]; | 
| 923 | 588 | |
| 5069 | 589 | Goal "(UN x:A. B(x)) = Union({Y. ? x:A. Y=B(x)})";
 | 
| 2891 | 590 | by (Blast_tac 1); | 
| 923 | 591 | qed "UN_eq"; | 
| 592 | ||
| 593 | (*Look: it has an EXISTENTIAL quantifier*) | |
| 5069 | 594 | Goal "(INT x:A. B(x)) = Inter({Y. ? x:A. Y=B(x)})";
 | 
| 2891 | 595 | by (Blast_tac 1); | 
| 923 | 596 | qed "INT_eq"; | 
| 597 | ||
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 598 | |
| 923 | 599 | (*Distributive laws...*) | 
| 600 | ||
| 5069 | 601 | Goal "A Int Union(B) = (UN C:B. A Int C)"; | 
| 2891 | 602 | by (Blast_tac 1); | 
| 923 | 603 | qed "Int_Union"; | 
| 604 | ||
| 5069 | 605 | Goal "Union(B) Int A = (UN C:B. C Int A)"; | 
| 4674 | 606 | by (Blast_tac 1); | 
| 607 | qed "Int_Union2"; | |
| 608 | ||
| 4306 
ddbe1a9722ab
Tidying and using equalityCE instead of the slower equalityE
 paulson parents: 
4231diff
changeset | 609 | (* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: | 
| 923 | 610 | Union of a family of unions **) | 
| 5069 | 611 | Goal "(UN x:C. A(x) Un B(x)) = Union(A``C) Un Union(B``C)"; | 
| 2891 | 612 | by (Blast_tac 1); | 
| 923 | 613 | qed "Un_Union_image"; | 
| 614 | ||
| 615 | (*Equivalent version*) | |
| 5069 | 616 | Goal "(UN i:I. A(i) Un B(i)) = (UN i:I. A(i)) Un (UN i:I. B(i))"; | 
| 2891 | 617 | by (Blast_tac 1); | 
| 923 | 618 | qed "UN_Un_distrib"; | 
| 619 | ||
| 5069 | 620 | Goal "A Un Inter(B) = (INT C:B. A Un C)"; | 
| 2891 | 621 | by (Blast_tac 1); | 
| 923 | 622 | qed "Un_Inter"; | 
| 623 | ||
| 5069 | 624 | Goal "(INT x:C. A(x) Int B(x)) = Inter(A``C) Int Inter(B``C)"; | 
| 2891 | 625 | by (Blast_tac 1); | 
| 923 | 626 | qed "Int_Inter_image"; | 
| 627 | ||
| 628 | (*Equivalent version*) | |
| 5069 | 629 | Goal "(INT i:I. A(i) Int B(i)) = (INT i:I. A(i)) Int (INT i:I. B(i))"; | 
| 2891 | 630 | by (Blast_tac 1); | 
| 923 | 631 | qed "INT_Int_distrib"; | 
| 632 | ||
| 633 | (*Halmos, Naive Set Theory, page 35.*) | |
| 5069 | 634 | Goal "B Int (UN i:I. A(i)) = (UN i:I. B Int A(i))"; | 
| 2891 | 635 | by (Blast_tac 1); | 
| 923 | 636 | qed "Int_UN_distrib"; | 
| 637 | ||
| 5069 | 638 | Goal "B Un (INT i:I. A(i)) = (INT i:I. B Un A(i))"; | 
| 2891 | 639 | by (Blast_tac 1); | 
| 923 | 640 | qed "Un_INT_distrib"; | 
| 641 | ||
| 5278 | 642 | Goal "(UN i:I. A(i)) Int (UN j:J. B(j)) = (UN i:I. UN j:J. A(i) Int B(j))"; | 
| 2891 | 643 | by (Blast_tac 1); | 
| 923 | 644 | qed "Int_UN_distrib2"; | 
| 645 | ||
| 5278 | 646 | Goal "(INT i:I. A(i)) Un (INT j:J. B(j)) = (INT i:I. INT j:J. A(i) Un B(j))"; | 
| 2891 | 647 | by (Blast_tac 1); | 
| 923 | 648 | qed "Un_INT_distrib2"; | 
| 649 | ||
| 2512 | 650 | |
| 651 | section"Bounded quantifiers"; | |
| 652 | ||
| 3860 | 653 | (** The following are not added to the default simpset because | 
| 654 | (a) they duplicate the body and (b) there are no similar rules for Int. **) | |
| 2512 | 655 | |
| 5069 | 656 | Goal "(ALL x:A Un B. P x) = ((ALL x:A. P x) & (ALL x:B. P x))"; | 
| 2891 | 657 | by (Blast_tac 1); | 
| 2519 | 658 | qed "ball_Un"; | 
| 659 | ||
| 5069 | 660 | Goal "(EX x:A Un B. P x) = ((EX x:A. P x) | (EX x:B. P x))"; | 
| 2891 | 661 | by (Blast_tac 1); | 
| 2519 | 662 | qed "bex_Un"; | 
| 2512 | 663 | |
| 5069 | 664 | Goal "(ALL z: UNION A B. P z) = (ALL x:A. ALL z:B x. P z)"; | 
| 4771 | 665 | by (Blast_tac 1); | 
| 666 | qed "ball_UN"; | |
| 667 | ||
| 5069 | 668 | Goal "(EX z: UNION A B. P z) = (EX x:A. EX z:B x. P z)"; | 
| 4771 | 669 | by (Blast_tac 1); | 
| 670 | qed "bex_UN"; | |
| 671 | ||
| 2512 | 672 | |
| 1548 | 673 | section "-"; | 
| 923 | 674 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 675 | Goal "A-B = A Int (-B)"; | 
| 4609 | 676 | by (Blast_tac 1); | 
| 4662 | 677 | qed "Diff_eq"; | 
| 4609 | 678 | |
| 7516 | 679 | Goal "(A-B = {}) = (A<=B)";
 | 
| 680 | by (Blast_tac 1); | |
| 681 | qed "Diff_eq_empty_iff"; | |
| 682 | Addsimps[Diff_eq_empty_iff]; | |
| 683 | ||
| 5069 | 684 | Goal "A-A = {}";
 | 
| 2891 | 685 | by (Blast_tac 1); | 
| 923 | 686 | qed "Diff_cancel"; | 
| 1531 | 687 | Addsimps[Diff_cancel]; | 
| 923 | 688 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 689 | Goal "A Int B = {} ==> A-B = A";
 | 
| 4674 | 690 | by (blast_tac (claset() addEs [equalityE]) 1); | 
| 691 | qed "Diff_triv"; | |
| 692 | ||
| 5069 | 693 | Goal "{}-A = {}";
 | 
| 2891 | 694 | by (Blast_tac 1); | 
| 923 | 695 | qed "empty_Diff"; | 
| 1531 | 696 | Addsimps[empty_Diff]; | 
| 923 | 697 | |
| 5069 | 698 | Goal "A-{} = A";
 | 
| 2891 | 699 | by (Blast_tac 1); | 
| 923 | 700 | qed "Diff_empty"; | 
| 1531 | 701 | Addsimps[Diff_empty]; | 
| 702 | ||
| 5069 | 703 | Goal "A-UNIV = {}";
 | 
| 2891 | 704 | by (Blast_tac 1); | 
| 1531 | 705 | qed "Diff_UNIV"; | 
| 706 | Addsimps[Diff_UNIV]; | |
| 707 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 708 | Goal "x~:A ==> A - insert x B = A-B"; | 
| 2891 | 709 | by (Blast_tac 1); | 
| 1531 | 710 | qed "Diff_insert0"; | 
| 711 | Addsimps [Diff_insert0]; | |
| 923 | 712 | |
| 713 | (*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
 | |
| 5069 | 714 | Goal "A - insert a B = A - B - {a}";
 | 
| 2891 | 715 | by (Blast_tac 1); | 
| 923 | 716 | qed "Diff_insert"; | 
| 717 | ||
| 718 | (*NOT SUITABLE FOR REWRITING since {a} == insert a 0*)
 | |
| 5069 | 719 | Goal "A - insert a B = A - {a} - B";
 | 
| 2891 | 720 | by (Blast_tac 1); | 
| 923 | 721 | qed "Diff_insert2"; | 
| 722 | ||
| 5069 | 723 | Goal "insert x A - B = (if x:B then A-B else insert x (A-B))"; | 
| 4686 | 724 | by (Simp_tac 1); | 
| 2891 | 725 | by (Blast_tac 1); | 
| 1531 | 726 | qed "insert_Diff_if"; | 
| 727 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 728 | Goal "x:B ==> insert x A - B = A-B"; | 
| 2891 | 729 | by (Blast_tac 1); | 
| 1531 | 730 | qed "insert_Diff1"; | 
| 731 | Addsimps [insert_Diff1]; | |
| 732 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 733 | Goal "a:A ==> insert a (A-{a}) = A";
 | 
| 2922 | 734 | by (Blast_tac 1); | 
| 923 | 735 | qed "insert_Diff"; | 
| 736 | ||
| 7824 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 737 | Goal "x ~: A ==> (insert x A) - {x} = A";
 | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 738 | by Auto_tac; | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 739 | qed "Diff_insert_absorb"; | 
| 
1a85ba81d019
new default simprule Collect_const and new them Diff_insert_absorb
 paulson parents: 
7713diff
changeset | 740 | |
| 5069 | 741 | Goal "A Int (B-A) = {}";
 | 
| 2891 | 742 | by (Blast_tac 1); | 
| 923 | 743 | qed "Diff_disjoint"; | 
| 1531 | 744 | Addsimps[Diff_disjoint]; | 
| 923 | 745 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 746 | Goal "A<=B ==> A Un (B-A) = B"; | 
| 2891 | 747 | by (Blast_tac 1); | 
| 923 | 748 | qed "Diff_partition"; | 
| 749 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 750 | Goal "[| A<=B; B<= C |] ==> (B - (C - A)) = (A :: 'a set)"; | 
| 2891 | 751 | by (Blast_tac 1); | 
| 923 | 752 | qed "double_diff"; | 
| 753 | ||
| 5069 | 754 | Goal "A Un (B-A) = A Un B"; | 
| 4645 | 755 | by (Blast_tac 1); | 
| 756 | qed "Un_Diff_cancel"; | |
| 757 | ||
| 5069 | 758 | Goal "(B-A) Un A = B Un A"; | 
| 4645 | 759 | by (Blast_tac 1); | 
| 760 | qed "Un_Diff_cancel2"; | |
| 761 | ||
| 762 | Addsimps [Un_Diff_cancel, Un_Diff_cancel2]; | |
| 763 | ||
| 5069 | 764 | Goal "A - (B Un C) = (A-B) Int (A-C)"; | 
| 2891 | 765 | by (Blast_tac 1); | 
| 923 | 766 | qed "Diff_Un"; | 
| 767 | ||
| 5069 | 768 | Goal "A - (B Int C) = (A-B) Un (A-C)"; | 
| 2891 | 769 | by (Blast_tac 1); | 
| 923 | 770 | qed "Diff_Int"; | 
| 771 | ||
| 5069 | 772 | Goal "(A Un B) - C = (A - C) Un (B - C)"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 773 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 774 | qed "Un_Diff"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 775 | |
| 5069 | 776 | Goal "(A Int B) - C = A Int (B - C)"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 777 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 778 | qed "Int_Diff"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 779 | |
| 5069 | 780 | Goal "C Int (A-B) = (C Int A) - (C Int B)"; | 
| 4748 | 781 | by (Blast_tac 1); | 
| 782 | qed "Diff_Int_distrib"; | |
| 783 | ||
| 5069 | 784 | Goal "(A-B) Int C = (A Int C) - (B Int C)"; | 
| 4645 | 785 | by (Blast_tac 1); | 
| 4748 | 786 | qed "Diff_Int_distrib2"; | 
| 4645 | 787 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
6832diff
changeset | 788 | Goal "A - (- B) = A Int B"; | 
| 5632 | 789 | by Auto_tac; | 
| 790 | qed "Diff_Compl"; | |
| 791 | Addsimps [Diff_Compl]; | |
| 792 | ||
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 793 | |
| 5238 | 794 | section "Quantification over type \"bool\""; | 
| 795 | ||
| 796 | Goal "(ALL b::bool. P b) = (P True & P False)"; | |
| 797 | by Auto_tac; | |
| 798 | by (case_tac "b" 1); | |
| 799 | by Auto_tac; | |
| 800 | qed "all_bool_eq"; | |
| 801 | ||
| 5762 | 802 | bind_thm ("bool_induct", conjI RS (all_bool_eq RS iffD2) RS spec);
 | 
| 803 | ||
| 5238 | 804 | Goal "(EX b::bool. P b) = (P True | P False)"; | 
| 805 | by Auto_tac; | |
| 806 | by (case_tac "b" 1); | |
| 807 | by Auto_tac; | |
| 808 | qed "ex_bool_eq"; | |
| 809 | ||
| 810 | Goal "A Un B = (UN b. if b then A else B)"; | |
| 8026 | 811 | by (auto_tac(claset(), simpset() addsimps [split_if_mem2])); | 
| 5238 | 812 | qed "Un_eq_UN"; | 
| 813 | ||
| 814 | Goal "(UN b::bool. A b) = (A True Un A False)"; | |
| 815 | by Auto_tac; | |
| 816 | by (case_tac "b" 1); | |
| 817 | by Auto_tac; | |
| 818 | qed "UN_bool_eq"; | |
| 819 | ||
| 820 | Goal "(INT b::bool. A b) = (A True Int A False)"; | |
| 821 | by Auto_tac; | |
| 822 | by (case_tac "b" 1); | |
| 823 | by Auto_tac; | |
| 824 | qed "INT_bool_eq"; | |
| 825 | ||
| 826 | ||
| 6292 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 827 | section "Pow"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 828 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 829 | Goalw [Pow_def] "Pow {} = {{}}";
 | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 830 | by Auto_tac; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 831 | qed "Pow_empty"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 832 | Addsimps [Pow_empty]; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 833 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 834 | Goal "Pow (insert a A) = Pow A Un (insert a `` Pow A)"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 835 | by Safe_tac; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 836 | by (etac swap 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 837 | by (res_inst_tac [("x", "x-{a}")] image_eqI 1);
 | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 838 | by (ALLGOALS Blast_tac); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 839 | qed "Pow_insert"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 840 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 841 | Goal "Pow (- A) = {-B |B. A: Pow B}";
 | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 842 | by Safe_tac; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 843 | by (Blast_tac 2); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 844 | by (res_inst_tac [("x", "-x")] exI 1);
 | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 845 | by (ALLGOALS Blast_tac); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 846 | qed "Pow_Compl"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 847 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 848 | Goal "Pow UNIV = UNIV"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 849 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 850 | qed "Pow_UNIV"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 851 | Addsimps [Pow_UNIV]; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 852 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 853 | Goal "Pow(A) Un Pow(B) <= Pow(A Un B)"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 854 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 855 | qed "Un_Pow_subset"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 856 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 857 | Goal "(UN x:A. Pow(B(x))) <= Pow(UN x:A. B(x))"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 858 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 859 | qed "UN_Pow_subset"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 860 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 861 | Goal "A <= Pow(Union(A))"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 862 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 863 | qed "subset_Pow_Union"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 864 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 865 | Goal "Union(Pow(A)) = A"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 866 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 867 | qed "Union_Pow_eq"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 868 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 869 | Goal "Pow(A Int B) = Pow(A) Int Pow(B)"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 870 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 871 | qed "Pow_Int_eq"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 872 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 873 | Goal "Pow(INT x:A. B(x)) = (INT x:A. Pow(B(x)))"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 874 | by (Blast_tac 1); | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 875 | qed "Pow_INT_eq"; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 876 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 877 | Addsimps [Union_Pow_eq, Pow_Int_eq]; | 
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 878 | |
| 
e50e1142dd06
new results e.g. about Pow; new simprules Union_image_eq, Inter_image_eq
 paulson parents: 
6283diff
changeset | 879 | |
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 880 | section "Miscellany"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 881 | |
| 5069 | 882 | Goal "(A = B) = ((A <= (B::'a set)) & (B<=A))"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 883 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 884 | qed "set_eq_subset"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 885 | |
| 5069 | 886 | Goal "A <= B = (! t. t:A --> t:B)"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 887 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 888 | qed "subset_iff"; | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 889 | |
| 5069 | 890 | Goalw [psubset_def] "((A::'a set) <= B) = ((A < B) | (A=B))"; | 
| 3222 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 891 | by (Blast_tac 1); | 
| 
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
 nipkow parents: 
2922diff
changeset | 892 | qed "subset_iff_psubset_eq"; | 
| 2021 | 893 | |
| 5069 | 894 | Goal "(!x. x ~: A) = (A={})";
 | 
| 4423 | 895 | by (Blast_tac 1); | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 896 | qed "all_not_in_conv"; | 
| 3907 | 897 | AddIffs [all_not_in_conv]; | 
| 3896 
ee8ebb74ec00
Various new lemmas. Improved conversion of equations to rewrite rules:
 nipkow parents: 
3860diff
changeset | 898 | |
| 6007 | 899 | |
| 5189 
362e4d6213c5
Added theorem distinct_lemma (needed for datatypes).
 berghofe parents: 
5148diff
changeset | 900 | (** for datatypes **) | 
| 
362e4d6213c5
Added theorem distinct_lemma (needed for datatypes).
 berghofe parents: 
5148diff
changeset | 901 | Goal "f x ~= f y ==> x ~= y"; | 
| 
362e4d6213c5
Added theorem distinct_lemma (needed for datatypes).
 berghofe parents: 
5148diff
changeset | 902 | by (Fast_tac 1); | 
| 
362e4d6213c5
Added theorem distinct_lemma (needed for datatypes).
 berghofe parents: 
5148diff
changeset | 903 | qed "distinct_lemma"; | 
| 
362e4d6213c5
Added theorem distinct_lemma (needed for datatypes).
 berghofe parents: 
5148diff
changeset | 904 | |
| 2021 | 905 | |
| 906 | (** Miniscoping: pushing in big Unions and Intersections **) | |
| 907 | local | |
| 4059 | 908 | fun prover s = prove_goal thy s (fn _ => [Blast_tac 1]) | 
| 2021 | 909 | in | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 910 | val UN_simps = map prover | 
| 5941 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 911 | ["!!C. c: C ==> (UN x:C. insert a (B x)) = insert a (UN x:C. B x)", | 
| 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 912 | "!!C. c: C ==> (UN x:C. A x Un B) = ((UN x:C. A x) Un B)", | 
| 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 913 | "!!C. c: C ==> (UN x:C. A Un B x) = (A Un (UN x:C. B x))", | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4136diff
changeset | 914 | "(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: 
4136diff
changeset | 915 | "(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: 
4136diff
changeset | 916 | "(UN x:C. A x - B) = ((UN x:C. A x) - B)", | 
| 4231 | 917 | "(UN x:C. A - B x) = (A - (INT x:C. B x))", | 
| 7914 | 918 | "(UN x: Union A. B x) = (UN y:A. UN x:y. B x)", | 
| 919 | "(UN z: UNION A B. C z) = (UN x:A. UN z: B(x). C z)", | |
| 4231 | 920 | "(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: 
2512diff
changeset | 921 | |
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 922 | val INT_simps = map prover | 
| 5941 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 923 | ["!!C. c: C ==> (INT x:C. A x Int B) = ((INT x:C. A x) Int B)", | 
| 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 924 | "!!C. c: C ==> (INT x:C. A Int B x) = (A Int (INT x:C. B x))", | 
| 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 925 | "!!C. c: C ==> (INT x:C. A x - B) = ((INT x:C. A x) - B)", | 
| 
1db9fad40a4f
better miniscoping rules: the premise C~={} is not good
 paulson parents: 
5931diff
changeset | 926 | "!!C. c: C ==> (INT x:C. A - B x) = (A - (UN x:C. B x))", | 
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4136diff
changeset | 927 | "(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: 
4136diff
changeset | 928 | "(INT x:C. A x Un B) = ((INT x:C. A x) Un B)", | 
| 4231 | 929 | "(INT x:C. A Un B x) = (A Un (INT x:C. B x))", | 
| 7914 | 930 | "(INT x: Union A. B x) = (INT y:A. INT x:y. B x)", | 
| 931 | "(INT z: UNION A B. C z) = (INT x:A. INT z: B(x). C z)", | |
| 4231 | 932 | "(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: 
2512diff
changeset | 933 | |
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 934 | |
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 935 | val ball_simps = map prover | 
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 936 | ["(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: 
2512diff
changeset | 937 | "(ALL x:A. P | Q x) = (P | (ALL x:A. Q x))", | 
| 3422 | 938 | "(ALL x:A. P --> Q x) = (P --> (ALL x:A. Q x))", | 
| 939 | "(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: 
2512diff
changeset | 940 |      "(ALL x:{}. P x) = True",
 | 
| 4136 | 941 | "(ALL x:UNIV. P x) = (ALL x. P x)", | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 942 | "(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: 
2512diff
changeset | 943 | "(ALL x:Union(A). P x) = (ALL y:A. ALL x:y. P x)", | 
| 5233 
3571ff68ceda
New rewrite rules for quantification over bounded UNIONs
 paulson parents: 
5189diff
changeset | 944 | "(ALL x: UNION A B. P x) = (ALL a:A. ALL x: B a. P x)", | 
| 3860 | 945 | "(ALL x:Collect Q. P x) = (ALL x. Q x --> P x)", | 
| 946 | "(ALL x:f``A. P x) = (ALL x:A. P(f x))", | |
| 947 | "(~(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: 
2512diff
changeset | 948 | |
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 949 | val ball_conj_distrib = | 
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 950 | 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: 
2512diff
changeset | 951 | |
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 952 | val bex_simps = map prover | 
| 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 953 | ["(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: 
2512diff
changeset | 954 | "(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: 
2512diff
changeset | 955 |      "(EX x:{}. P x) = False",
 | 
| 4136 | 956 | "(EX x:UNIV. P x) = (EX x. P x)", | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 957 | "(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: 
2512diff
changeset | 958 | "(EX x:Union(A). P x) = (EX y:A. EX x:y. P x)", | 
| 5233 
3571ff68ceda
New rewrite rules for quantification over bounded UNIONs
 paulson parents: 
5189diff
changeset | 959 | "(EX x: UNION A B. P x) = (EX a:A. EX x: B a. P x)", | 
| 3860 | 960 | "(EX x:Collect Q. P x) = (EX x. Q x & P x)", | 
| 961 | "(EX x:f``A. P x) = (EX x:A. P(f x))", | |
| 962 | "(~(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: 
2512diff
changeset | 963 | |
| 3426 | 964 | val bex_disj_distrib = | 
| 2513 
d708d8cdc8e8
New miniscoping rules for the bounded quantifiers and UN/INT operators
 paulson parents: 
2512diff
changeset | 965 | 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: 
2512diff
changeset | 966 | |
| 2021 | 967 | end; | 
| 968 | ||
| 7648 | 969 | bind_thms ("UN_simps", UN_simps);
 | 
| 970 | bind_thms ("INT_simps", INT_simps);
 | |
| 971 | bind_thms ("ball_simps", ball_simps);
 | |
| 972 | bind_thms ("bex_simps", bex_simps);
 | |
| 973 | bind_thm ("ball_conj_distrib", ball_conj_distrib);
 | |
| 974 | bind_thm ("bex_disj_distrib", bex_disj_distrib);
 | |
| 975 | ||
| 4159 
4aff9b7e5597
UNIV now a constant; UNION1, INTER1 now translations and no longer have
 paulson parents: 
4136diff
changeset | 976 | Addsimps (UN_simps @ INT_simps @ ball_simps @ bex_simps); |