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